Nominal/Ex/Weakening.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 25 Jun 2011 21:28:24 +0100
changeset 2900 d66430c7c4f1
parent 2645 09cf78bb53d4
child 2950 0911cb7bf696
permissions -rw-r--r--
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2645
Christian Urban <urbanc@in.tum.de>
parents: 2643
diff changeset
     1
theory Weakening
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
     5
section {* The Weakening property in the simply-typed lambda-calculus *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
atom_decl name
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
nominal_datatype lam =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  Var "name"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
| App "lam" "lam"
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    12
| Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100,100] 100)
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    14
text {* Typing *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
nominal_datatype ty =
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  TVar string
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    18
| TFun ty ty ("_ \<rightarrow> _" [100,100] 100) 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    20
lemma fresh_ty:
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  fixes x::"name"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  and   T::"ty"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  shows "atom x \<sharp> T"
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    24
  by (nominal_induct T rule: ty.strong_induct)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    25
     (simp_all add: ty.fresh pure_fresh)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    26
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    27
text {* Valid typing contexts *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
inductive
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  valid :: "(name \<times> ty) list \<Rightarrow> bool"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
where
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    32
  v_Nil[intro]: "valid []"
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    33
| v_Cons[intro]: "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T) # Gamma)"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    34
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    35
equivariance valid
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    36
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    37
text {* Typing judgements *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
inductive
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    40
  typing :: "(name \<times> ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
where
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    44
  | t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. t : T1 \<rightarrow> T2"
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    46
equivariance typing
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    48
text {* Strong induction principle for typing judgements *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
nominal_inductive typing
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  avoids t_Lam: "x"
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    52
  by (simp_all add: fresh_star_def fresh_ty lam.fresh)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
    53
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    55
abbreviation
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    56
  "sub_context" :: "(name \<times> ty) list \<Rightarrow> (name \<times> ty) list \<Rightarrow> bool" (infixr "\<subseteq>" 60) 
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    57
where
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    58
  "\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x, T) \<in> set \<Gamma>1 \<longrightarrow> (x, T) \<in> set \<Gamma>2"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    59
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    60
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    61
text {* The proof *}
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    63
lemma weakening_version1: 
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    64
  fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    65
  and   t ::"lam"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    66
  and   \<tau> ::"ty"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    67
  assumes a: "\<Gamma>1 \<turnstile> t : T"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    68
  and     b: "valid \<Gamma>2" 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    69
  and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    70
  shows "\<Gamma>2 \<turnstile> t : T"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    71
using a b c
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    72
proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    73
  case (t_Var \<Gamma>1 x T)  (* variable case *)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    74
  have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    75
  moreover  
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    76
  have "valid \<Gamma>2" by fact 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    77
  moreover 
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    78
  have "(x, T) \<in> set \<Gamma>1" by fact
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    79
  ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    80
next
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    81
  case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    82
  have vc: "atom x \<sharp> \<Gamma>2" by fact   (* variable convention *)
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    83
  have ih: "\<lbrakk>valid ((x, T1) # \<Gamma>2); (x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2\<rbrakk> \<Longrightarrow> (x, T1) # \<Gamma>2 \<turnstile> t : T2" by fact
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    84
  have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    85
  then have "(x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2" by simp
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    86
  moreover
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    87
  have "valid \<Gamma>2" by fact
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    88
  then have "valid ((x, T1) # \<Gamma>2)" using vc by auto
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    89
  ultimately have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    90
  with vc show "\<Gamma>2 \<turnstile> Lam [x]. t : T1 \<rightarrow> T2" by auto
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    91
qed (auto) (* app case *)
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
    93
lemma weakening_version2: 
2638
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    94
  fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    95
  assumes a: "\<Gamma>1 \<turnstile> t : T" 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    96
  and     b: "valid \<Gamma>2" 
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    97
  and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    98
  shows "\<Gamma>2 \<turnstile> t : T"
e1e2ca92760b strong rule inductions; as an example the weakening lemma works
Christian Urban <urbanc@in.tum.de>
parents: 2637
diff changeset
    99
using a b c
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   100
by (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   101
   (auto | atomize)+
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
   102
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
   103
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   104
text {* A version with finite sets as typing contexts *}
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   105
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   106
inductive
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   107
  valid2 :: "(name \<times> ty) fset \<Rightarrow> bool"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   108
where
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   109
  v2_Empty[intro]: "valid2 {||}"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   110
| v2_Insert[intro]: "\<lbrakk>(x, T) |\<notin>| Gamma; valid2 Gamma\<rbrakk> \<Longrightarrow> valid2 (insert_fset (x, T) Gamma)"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   111
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   112
equivariance valid2
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   113
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   114
inductive
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   115
  typing2 :: "(name \<times> ty) fset \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ 2\<turnstile> _ : _" [60,60,60] 60) 
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   116
where
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   117
    t2_Var[intro]: "\<lbrakk>valid2 \<Gamma>; (x, T) |\<in>| \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> 2\<turnstile> Var x : T"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   118
  | t2_App[intro]: "\<lbrakk>\<Gamma> 2\<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> 2\<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> 2\<turnstile> App t1 t2 : T2"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   119
  | t2_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; insert_fset (x, T1) \<Gamma> 2\<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> 2\<turnstile> Lam [x]. t : T1 \<rightarrow> T2"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   120
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   121
equivariance typing2
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   122
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   123
nominal_inductive typing2
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   124
  avoids t2_Lam: "x"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   125
  by (simp_all add: fresh_star_def fresh_ty lam.fresh)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   126
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   127
lemma weakening_version3: 
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   128
  fixes \<Gamma>::"(name \<times> ty) fset"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   129
  assumes a: "\<Gamma> 2\<turnstile> t : T" 
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   130
  and     b: "(x, T') |\<notin>| \<Gamma>"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   131
  shows "{|(x, T')|} |\<union>| \<Gamma> 2\<turnstile> t : T"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   132
using a b
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   133
apply(nominal_induct \<Gamma> t T avoiding: x rule: typing2.strong_induct)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   134
apply(auto)[2]
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   135
apply(rule t2_Lam)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   136
apply(simp add: fresh_insert_fset fresh_Pair fresh_ty)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
   137
apply(simp)
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   138
apply(drule_tac x="xa" in meta_spec)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
   139
apply(drule meta_mp)
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   140
apply(simp add: fresh_at_base)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   141
apply(simp add: insert_fset_left_comm)
2639
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
   142
done
a8fc346deda3 exported the code into a separate file
Christian Urban <urbanc@in.tum.de>
parents: 2638
diff changeset
   143
2643
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   144
lemma weakening_version4: 
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   145
  fixes \<Gamma>::"(name \<times> ty) fset"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   146
  assumes a: "\<Gamma> 2\<turnstile> t : T" 
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   147
  and     b: "(x, T') |\<notin>| \<Gamma>"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   148
  shows "{|(x, T')|} |\<union>| \<Gamma> 2\<turnstile> t : T"
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   149
using a b
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   150
apply(induct \<Gamma> t T arbitrary: x)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   151
apply(auto)[2]
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   152
apply(rule t2_Lam)
0579d3a48304 cleaned up weakening proof and added a version with finit sets
Christian Urban <urbanc@in.tum.de>
parents: 2639
diff changeset
   153
oops
2636
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
end
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
0865caafbfe6 file with most of the strong rule induction development
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158