Tutorial/Tutorial2.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 21 Jan 2011 22:58:03 +0100
changeset 2692 da9bed7baf23
parent 2689 ddc05a611005
child 2695 e8736c1cdd7f
permissions -rw-r--r--
better flow of proofs and definitions and proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2689
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Tutorial2
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Tutorial1
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
(* fixme: put height example here *)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
section {* Types *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
nominal_datatype ty =
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  tVar "string"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
| tArr "ty" "ty" ("_ \<rightarrow> _" [100, 100] 100)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
text {* 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  Having defined them as nominal datatypes gives us additional
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  definitions and theorems that come with types (see below).
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  We next define the type of typing contexts, a predicate that
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  defines valid contexts (i.e. lists that contain only unique
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  variables) and the typing judgement.
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
type_synonym ty_ctx = "(name \<times> ty) list"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
inductive
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  valid :: "ty_ctx \<Rightarrow> bool"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
where
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  v1[intro]: "valid []"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
| v2[intro]: "\<lbrakk>valid \<Gamma>; atom x \<sharp> \<Gamma>\<rbrakk>\<Longrightarrow> valid ((x, T) # \<Gamma>)"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
inductive
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  typing :: "ty_ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60, 60, 60] 60) 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
where
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  t_Var[intro]:  "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
| t_App[intro]:  "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
| t_Lam[intro]:  "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1 \<rightarrow> T2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
text {*
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  The predicate atom x \<sharp> \<Gamma>, read as x fresh for \<Gamma>, is defined by 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  Nominal Isabelle. It is defined for lambda-terms, products, lists etc. 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  For example we have:
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
lemma
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  fixes x::"name"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  shows "atom x \<sharp> Lam [x].t"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  and   "atom x \<sharp> (t1, t2) \<Longrightarrow> atom x \<sharp> App t1 t2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  and   "atom x \<sharp> Var y \<Longrightarrow> atom x \<sharp> y" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  and   "\<lbrakk>atom x \<sharp> t1; atom x \<sharp> t2\<rbrakk> \<Longrightarrow> atom x \<sharp> (t1, t2)"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  and   "\<lbrakk>atom x \<sharp> l1; atom x \<sharp> l2\<rbrakk> \<Longrightarrow> atom x \<sharp> (l1 @ l2)"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  and   "atom x \<sharp> y \<Longrightarrow> x \<noteq> y"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  by (simp_all add: lam.fresh fresh_append fresh_at_base) 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
text {* 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  We can also prove that every variable is fresh for a type. 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
lemma ty_fresh:
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  fixes x::"name"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  and   T::"ty"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  shows "atom x \<sharp> T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
by (induct T rule: ty.induct)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
   (simp_all add: ty.fresh pure_fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
text {* 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  In order to state the weakening lemma in a convenient form, we 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  using the following abbreviation. Abbreviations behave like 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  definitions, except that they are always automatically folded 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  and unfolded.
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
abbreviation
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  "sub_ty_ctx" :: "ty_ctx \<Rightarrow> ty_ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _" [60, 60] 60) 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
where
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  "\<Gamma>1 \<sqsubseteq> \<Gamma>2 \<equiv> \<forall>x. x \<in> set \<Gamma>1 \<longrightarrow> x \<in> set \<Gamma>2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
subsection {* EXERCISE 4 *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
text {*
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  Fill in the details and give a proof of the weakening lemma. 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
lemma 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
  assumes a: "\<Gamma>1 \<turnstile> t : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
  and     b: "valid \<Gamma>2" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
  shows "\<Gamma>2 \<turnstile> t : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
using a b c
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
proof (induct arbitrary: \<Gamma>2)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  case (t_Var \<Gamma>1 x T)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  have a1: "valid \<Gamma>1" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  have a2: "(x, T) \<in> set \<Gamma>1" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  have a3: "valid \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  have a4: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  show "\<Gamma>2 \<turnstile> Var x : T" sorry
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
next
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  case (t_Lam x \<Gamma>1 T1 t T2) 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  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
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  have a0: "atom x \<sharp> \<Gamma>1" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  have a1: "valid \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
qed (auto) -- {* the application case *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
text {* 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  Despite the frequent claim that the weakening lemma is trivial,
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  routine or obvious, the proof in the lambda-case does not go 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  through smoothly. Painful variable renamings seem to be necessary. 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  But the proof using renamings is horribly complicated (see below). 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
equivariance valid
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
equivariance typing
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
lemma weakening_NOT_TO_BE_TRIED_AT_HOME: 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  assumes a: "\<Gamma>1 \<turnstile> t : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  and     b: "valid \<Gamma>2" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  shows "\<Gamma>2 \<turnstile> t : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
using a b c
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
proof (induct arbitrary: \<Gamma>2)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  -- {* lambda case *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  case (t_Lam x \<Gamma>1 T1 t T2) 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  have fc0: "atom x \<sharp> \<Gamma>1" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  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
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  -- {* we choose a fresh variable *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  obtain c::"name" where fc1: "atom c \<sharp> (x, t, \<Gamma>1, \<Gamma>2)" by (rule obtain_fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
  -- {* we alpha-rename the abstraction *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
  have "Lam [c].((c \<leftrightarrow> x) \<bullet> t) = Lam [x].t" using fc1
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
    by (auto simp add: lam.eq_iff Abs1_eq_iff flip_def)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  moreover
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  -- {* we can then alpha rename the goal *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  have "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
  proof - 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
    -- {* we need to establish *} 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
    -- {* (1)   (x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
    -- {* (2)   valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)) *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
    have "(1)": "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
    proof -
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
      have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
      then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc0 fc1
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
        by (perm_simp) (simp add: flip_fresh_fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
      then show "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2)" by (rule permute_boolE)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
    qed
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
    moreover 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
    have "(2)": "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
    proof -
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
      have "valid \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
      then show "valid ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2))" using fc1
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
        by (auto simp add: fresh_permute_left atom_eqvt valid.eqvt)	
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
    qed
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
    -- {* these two facts give us by induction hypothesis the following *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
    ultimately have "(x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2" using ih by simp 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
    -- {* we now apply renamings to get to our goal *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
    then have "(c \<leftrightarrow> x) \<bullet> ((x, T1) # ((c \<leftrightarrow> x) \<bullet> \<Gamma>2) \<turnstile> t : T2)" by (rule permute_boolI)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
    then have "(c, T1) # \<Gamma>2 \<turnstile> ((c \<leftrightarrow> x) \<bullet> t) : T2" using fc1
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
      by (perm_simp) (simp add: flip_fresh_fresh ty_fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
    then show "\<Gamma>2 \<turnstile> Lam [c].((c \<leftrightarrow> x) \<bullet> t) : T1 \<rightarrow> T2" using fc1 by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  qed
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
qed (auto) -- {* var and app cases, luckily, are automatic *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
text {* 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  The remedy is to use a stronger induction principle that
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  has the usual "variable convention" already build in. The
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  following command derives this induction principle for us.
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  (We shall explain what happens here later.)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
nominal_inductive typing
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
  avoids t_Lam: "x"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
  unfolding fresh_star_def
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
  by (simp_all add: fresh_Pair lam.fresh ty_fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
text {* Compare the two induction principles *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
thm typing.induct
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
thm typing.strong_induct -- {* has the additional assumption {atom x} \<sharp> c *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
text {* 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  We can use the stronger induction principle by replacing
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  the line
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
      proof (induct arbitrary: \<Gamma>2)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  with 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
      proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  Try now the proof.
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
*}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
lemma weakening: 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  assumes a: "\<Gamma>1 \<turnstile> t : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
  and     b: "valid \<Gamma>2" 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  and     c: "\<Gamma>1 \<sqsubseteq> \<Gamma>2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  shows "\<Gamma>2 \<turnstile> t : T"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
using a b c
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  case (t_Var \<Gamma>1 x T)  -- {* variable case is as before *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  have "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
  moreover  
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  have "valid \<Gamma>2" by fact 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  moreover 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  have "(x, T)\<in> set \<Gamma>1" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
next
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  case (t_Lam x \<Gamma>1 T1 t T2) 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  have vc: "atom x \<sharp> \<Gamma>2" by fact  -- {* additional fact afforded by the strong induction *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  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
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
  have a0: "atom x \<sharp> \<Gamma>1" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  have a1: "valid \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  have a2: "\<Gamma>1 \<sqsubseteq> \<Gamma>2" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
  have "valid ((x, T1) # \<Gamma>2)" using a1 vc by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
  moreover
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  have "(x, T1) # \<Gamma>1 \<sqsubseteq> (x, T1) # \<Gamma>2" using a2 by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  ultimately 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp 
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  then show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" using vc by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
qed (auto) -- {* app case *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
text {* unbind / inconsistency example *}
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
inductive
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  unbind :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<mapsto> _" [60, 60] 60)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
where
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  u_Var[intro]: "Var x \<mapsto> Var x"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
| u_App[intro]: "App t1 t2 \<mapsto> App t1 t2"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
| u_Lam[intro]: "t \<mapsto> t' \<Longrightarrow> Lam [x].t \<mapsto> t'"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
lemma unbind_simple:
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
  shows "Lam [x].Var x \<mapsto> Var x"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
equivariance unbind
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
nominal_inductive unbind
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  avoids u_Lam: "x"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  sorry
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
lemma unbind_fake:
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  fixes y::"name"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  assumes a: "t \<mapsto> t'"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  and     b: "atom y \<sharp> t"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  shows "atom y \<sharp> t'"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
using a b
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
proof (nominal_induct avoiding: y rule: unbind.strong_induct)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  case (u_Lam t t' x y)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  have ih: "atom y \<sharp> t \<Longrightarrow> atom y \<sharp> t'" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
  have asm: "atom y \<sharp> Lam [x]. t" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  have fc: "atom x \<sharp> y" by fact
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  then have in_eq: "x \<noteq> y" by (simp add: fresh_at_base)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  then have "atom y \<sharp> t" using asm by (simp add: lam.fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
  then show "atom y \<sharp> t'" using ih by simp
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
qed
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
lemma
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  shows "False"
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
proof -
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  have "atom x \<sharp> Lam [x]. Var x" by (simp add: lam.fresh)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  then have "atom x \<sharp> Var x" using unbind_fake unbind_simple by auto
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  then show "False" by (simp add: lam.fresh fresh_at_base)
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
qed
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
ddc05a611005 added unbind example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
end