Tutorial/Tutorial2s.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 06 Jul 2011 01:04:09 +0200
changeset 2954 dc6007dd9c30
parent 2695 e8736c1cdd7f
child 3132 87eca760dcba
permissions -rw-r--r--
a little further with NBE
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2695
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Tutorial2s
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
imports Lambda
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
section {* Height of a Lambda-Term *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  The theory Lambda defines the notions of capture-avoiding
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  substitution and the height of lambda terms. 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
thm height.simps
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
thm subst.simps
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
subsection {* EXERCISE 7 *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  Lets prove a property about the height of substitutions.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  Assume that the height of a lambda-term is always greater 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  or equal to 1.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
lemma height_ge_one: 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  shows "1 \<le> height t"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
by (induct t rule: lam.induct) (auto)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
text {* Remove the sorries *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
theorem
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  shows "height (t[x ::= t']) \<le> height t - 1 + height t'"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
proof (nominal_induct t avoiding: x t' rule: lam.strong_induct)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
  case (Var y)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  have "1 \<le> height t'" using height_ge_one by simp
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  then show "height (Var y[x ::= t']) \<le> height (Var y) - 1 + height t'" by simp
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
next    
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  case (App t1 t2)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  have ih1: "height (t1[x::=t']) \<le> (height t1) - 1 + height t'" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  and  ih2: "height (t2[x::=t']) \<le> (height t2) - 1 + height t'" by fact+
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  then show "height ((App t1 t2)[x ::= t']) \<le> height (App t1 t2) - 1 + height t'" by simp  
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
next
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  case (Lam y t1)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  have ih: "height (t1[x::=t']) \<le> height t1 - 1 + height t'" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  moreover
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  have vc: "atom y \<sharp> x" "atom y \<sharp> t'" by fact+ -- {* usual variable convention *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  ultimately 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  show "height ((Lam [y].t1)[x ::= t']) \<le> height (Lam [y].t1) - 1 + height t'" by simp
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
qed
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
section {* Types and the Weakening Lemma *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
nominal_datatype ty =
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  tVar "string"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
| tArr "ty" "ty" (infixr "\<rightarrow>" 100)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  Having defined them as nominal datatypes gives us additional
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  definitions and theorems that come with types (see below).
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  We next define the type of typing contexts, a predicate that
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  defines valid contexts (i.e. lists that contain only unique
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  variables) and the typing judgement.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
type_synonym ty_ctx = "(name \<times> ty) list"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
inductive
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  valid :: "ty_ctx \<Rightarrow> bool"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
where
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  v1[intro]: "valid []"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
| v2[intro]: "\<lbrakk>valid \<Gamma>; atom x \<sharp> \<Gamma>\<rbrakk>\<Longrightarrow> valid ((x, T) # \<Gamma>)"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
inductive
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  typing :: "ty_ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60, 60, 60] 60) 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
where
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  t_Var[intro]:  "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
| t_App[intro]:  "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
| t_Lam[intro]:  "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1 \<rightarrow> T2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  The predicate atom x \<sharp> \<Gamma>, read as x fresh for \<Gamma>, is defined by 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  Nominal Isabelle. It is defined for lambda-terms, products, lists etc. 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  For example we have:
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
lemma
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
  fixes x::"name"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  shows "atom x \<sharp> Lam [x].t"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  and   "atom x \<sharp> (t1, t2) \<Longrightarrow> atom x \<sharp> App t1 t2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  and   "atom x \<sharp> Var y \<Longrightarrow> atom x \<sharp> y" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  and   "\<lbrakk>atom x \<sharp> t1; atom x \<sharp> t2\<rbrakk> \<Longrightarrow> atom x \<sharp> (t1, t2)"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  and   "\<lbrakk>atom x \<sharp> l1; atom x \<sharp> l2\<rbrakk> \<Longrightarrow> atom x \<sharp> (l1 @ l2)"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  and   "atom x \<sharp> y \<Longrightarrow> x \<noteq> y"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  by (simp_all add: lam.fresh fresh_append fresh_at_base) 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
  We can also prove that every variable is fresh for a type. 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
lemma ty_fresh:
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  fixes x::"name"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  and   T::"ty"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  shows "atom x \<sharp> T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
by (induct T rule: ty.induct)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
   (simp_all add: ty.fresh pure_fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  In order to state the weakening lemma in a convenient form, we 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  using the following abbreviation. Abbreviations behave like 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  definitions, except that they are always automatically folded 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  and unfolded.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
abbreviation
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  "sub_ty_ctx" :: "ty_ctx \<Rightarrow> ty_ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _" [60, 60] 60) 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
where
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  "\<Gamma>1 \<sqsubseteq> \<Gamma>2 \<equiv> \<forall>x. x \<in> set \<Gamma>1 \<longrightarrow> x \<in> set \<Gamma>2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
subsection {* EXERCISE 8 *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  Fill in the details and give a proof of the weakening lemma. 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
lemma 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  assumes a: "\<Gamma>1 \<turnstile> t : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  and     b: "valid \<Gamma>2" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  shows "\<Gamma>2 \<turnstile> t : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
using a b c
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
proof (induct arbitrary: \<Gamma>2)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
  case (t_Var \<Gamma>1 x T)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  have a1: "valid \<Gamma>1" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  have a2: "(x, T) \<in> set \<Gamma>1" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  have a3: "valid \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
  have a4: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  show "\<Gamma>2 \<turnstile> Var x : T" sorry
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
next
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  case (t_Lam x \<Gamma>1 T1 t T2) 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
  have a0: "atom x \<sharp> \<Gamma>1" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  have a1: "valid \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
qed (auto) -- {* the application case is automatic*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  Despite the frequent claim that the weakening lemma is trivial,
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  routine or obvious, the proof in the lambda-case does not go 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
  through smoothly. Painful variable renamings seem to be necessary. 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  But the proof using renamings is horribly complicated (see below). 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
equivariance valid
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
equivariance typing
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
lemma weakening_NOT_TO_BE_TRIED_AT_HOME: 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  assumes a: "\<Gamma>1 \<turnstile> t : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  and     b: "valid \<Gamma>2" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  shows "\<Gamma>2 \<turnstile> t : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
using a b c
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
proof (induct arbitrary: \<Gamma>2)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  -- {* lambda case *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  case (t_Lam x \<Gamma>1 T1 t T2) 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  have fc0: "atom x \<sharp> \<Gamma>1" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  -- {* we choose a fresh variable *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  obtain c::"name" where fc1: "atom c \<sharp> (x, t, \<Gamma>1, \<Gamma>2)" by (rule obtain_fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
  -- {* we alpha-rename the abstraction *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
  have "Lam [c].((c \<leftrightarrow> x) \<bullet> t) = Lam [x].t" using fc1
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
    by (auto simp add: lam.eq_iff Abs1_eq_iff flip_def)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
  moreover
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
  -- {* we can then alpha rename the goal *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  have "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  proof - 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
    -- {* we need to establish *} 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
    -- {* (1)   (x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
    -- {* (2)   valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)) *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
    have "(1)": "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
    proof -
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
      have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
      then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc0 fc1
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
        by (perm_simp) (simp add: flip_fresh_fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
      then show "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" by (rule permute_boolE)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
    qed
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
    moreover 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
    have "(2)": "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
    proof -
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
      have "valid \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
      then show "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc1
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
        by (auto simp add: fresh_permute_left atom_eqvt valid.eqvt)	
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
    qed
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
    -- {* these two facts give us by induction hypothesis the following *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
    ultimately have "(x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2" using ih by simp 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
    -- {* we now apply renamings to get to our goal *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
    then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2)" by (rule permute_boolI)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
    then have "(c, T1) # \<Gamma>2 \<turnstile> ((c \<leftrightarrow> x) \<bullet> t) : T2" using fc1
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
      by (perm_simp) (simp add: flip_fresh_fresh ty_fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
    then show "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" using fc1 by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  qed
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
qed (auto) -- {* var and app cases, luckily, are automatic *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  The remedy is to use again a stronger induction principle that
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
  has the usual "variable convention" already build in. The
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  following command derives this induction principle for the typing
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  relation. (We shall explain what happens here later.)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
nominal_inductive typing
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  avoids t_Lam: "x"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
  unfolding fresh_star_def
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
  by (simp_all add: fresh_Pair lam.fresh ty_fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
text {* Compare the two induction principles *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
thm typing.induct
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
thm typing.strong_induct -- {* note the additional assumption {atom x} \<sharp> c *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
text {* 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
  We can use the stronger induction principle by replacing
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  the line
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
      proof (induct arbitrary: \<Gamma>2)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  with 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
      proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  Try now the proof.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
subsection {* EXERCISE 9 *} 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
lemma weakening: 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  assumes a: "\<Gamma>1 \<turnstile> t : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  and     b: "valid \<Gamma>2" 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  shows "\<Gamma>2 \<turnstile> t : T"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
using a b c
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  case (t_Var \<Gamma>1 x T)  -- {* variable case is as before *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  moreover  
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  have "valid \<Gamma>2" by fact 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  moreover 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  have "(x, T)\<in> set \<Gamma>1" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
next
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  case (t_Lam x \<Gamma>1 T1 t T2) 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  have vc: "atom x \<sharp> \<Gamma>2" by fact  -- {* additional fact afforded by the strong induction *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x, T1) # \<Gamma>1 \<sqsubseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
  have a0: "atom x \<sharp> \<Gamma>1" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  have a1: "valid \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
  have "valid ((x, T1) # \<Gamma>2)" using a1 vc by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  moreover
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  have "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # \<Gamma>2" using a2 by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  ultimately 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  then show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" using vc by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
qed (auto) -- {* app case *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
section {* Unbind and an Inconsistency Example *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  You might wonder why we had to discharge some proof
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  obligations in order to obtain the stronger induction
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  principle for the typing relation. (Remember for
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
  lambda terms this stronger induction principle is
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  derived automatically.)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
  
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
  This proof obligation is really needed, because if we 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
  assume universally a stronger induction principle, then 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  in some cases we can derive false. This is "shown" below.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
inductive
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  unbind :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixr "\<mapsto>" 60)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
where
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
  u_Var[intro]: "Var x \<mapsto> Var x"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
| u_App[intro]: "App t1 t2 \<mapsto> App t1 t2"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
| u_Lam[intro]: "t \<mapsto> t' \<Longrightarrow> Lam [x].t \<mapsto> t'"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
text {* It is clear that the following judgement holds. *}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
lemma unbind_simple:
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
  shows "Lam [x].Var x \<mapsto> Var x"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  Now lets derive the strong induction principle for unbind.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  The point is that we cannot establish the proof obligations,
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
  therefore we force Isabelle to accept them by using "sorry".
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
equivariance unbind
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
nominal_inductive unbind
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
  avoids u_Lam: "x"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
  sorry
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
  Using the stronger induction principle, we can establish
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
  th follwoing bogus property.
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
lemma unbind_fake:
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
  fixes y::"name"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  assumes a: "t \<mapsto> t'"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
  and     b: "atom y \<sharp> t"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
  shows "atom y \<sharp> t'"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
using a b
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
proof (nominal_induct avoiding: y rule: unbind.strong_induct)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
  case (u_Lam t t' x y)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  have ih: "atom y \<sharp> t \<Longrightarrow> atom y \<sharp> t'" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  have asm: "atom y \<sharp> Lam [x]. t" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  have fc: "atom x \<sharp> y" by fact
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  then have in_eq: "x \<noteq> y" by (simp add: fresh_at_base)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
  then have "atom y \<sharp> t" using asm by (simp add: lam.fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
  then show "atom y \<sharp> t'" using ih by simp
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
qed
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
text {*
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
  And from this follows the inconsitency:
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
*}
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
lemma
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
  shows "False"
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
proof -
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
  have "atom x \<sharp> Lam [x]. Var x" by (simp add: lam.fresh)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
  moreover 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  have "Lam [x]. Var x \<mapsto> Var x" using unbind_simple by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
  ultimately 
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
  have "atom x \<sharp> Var x" using unbind_fake by auto
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
  then have "x \<noteq> x" by (simp add: lam.fresh fresh_at_base)
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
  then show "False" by simp
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
qed
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
e8736c1cdd7f cleaned up and solution section
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
end