Tutorial/Tutorial3.thy
changeset 2690 f325eefe803e
child 2691 abb6c3ac2df2
equal deleted inserted replaced
2689:ddc05a611005 2690:f325eefe803e
       
     1 theory Tutorial3
       
     2 imports Lambda
       
     3 begin
       
     4 
       
     5 section {* Formalising Barendregt's Proof of the Substitution Lemma *}
       
     6 
       
     7 text {*
       
     8   Barendregt's proof needs in the variable case a case distinction.
       
     9   One way to do this in Isar is to use blocks. A block is some sequent
       
    10   or reasoning steps enclosed in curly braces
       
    11 
       
    12   { \<dots>
       
    13 
       
    14     have "statement"
       
    15   }
       
    16 
       
    17   Such a block can contain local assumptions like
       
    18 
       
    19   { assume "A"
       
    20     assume "B"
       
    21     \<dots>
       
    22     have "C" by \<dots>
       
    23   }
       
    24 
       
    25   Where "C" is the last have-statement in this block. The behaviour 
       
    26   of such a block to the 'outside' is the implication
       
    27 
       
    28    \<lbrakk>A; B\<rbrakk> \<Longrightarrow> "C" 
       
    29 
       
    30   Now if we want to prove a property "smth" using the case-distinctions
       
    31   P1, P2 and P3 then we can use the following reasoning:
       
    32 
       
    33     { assume "P1"
       
    34       \<dots>
       
    35       have "smth"
       
    36     }
       
    37     moreover
       
    38     { assume "P2"
       
    39       \<dots>
       
    40       have "smth"
       
    41     }
       
    42     moreover
       
    43     { assume "P3"
       
    44       \<dots>
       
    45       have "smth"
       
    46     }
       
    47     ultimately have "smth" by blast
       
    48 
       
    49   The blocks establish the implications
       
    50 
       
    51     P1 \<Longrightarrow> smth
       
    52     P2 \<Longrightarrow> smth
       
    53     P3 \<Longrightarrow> smth
       
    54 
       
    55   If we know that P1, P2 and P3 cover all the cases, that is P1 \<or> P2 \<or> P3 is
       
    56   true, then we have 'ultimately' established the property "smth" 
       
    57   
       
    58 *}
       
    59 
       
    60 section {* EXERCISE 7 *}
       
    61 
       
    62 text {*
       
    63   Fill in the cases 1.2 and 1.3 and the equational reasoning 
       
    64   in the lambda-case.
       
    65 *}
       
    66 
       
    67 lemma forget:
       
    68   shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
       
    69 by (nominal_induct t avoiding: x s rule: lam.strong_induct)
       
    70    (auto simp add: lam.fresh fresh_at_base)
       
    71 
       
    72 lemma fresh_fact:
       
    73   assumes a: "atom z \<sharp> s"
       
    74   and b: "z = y \<or> atom z \<sharp> t"
       
    75   shows "atom z \<sharp> t[y ::= s]"
       
    76 using a b
       
    77 by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
       
    78    (auto simp add: lam.fresh fresh_at_base)
       
    79 
       
    80 
       
    81 lemma 
       
    82   assumes a: "x \<noteq> y"
       
    83   and     b: "atom x \<sharp> L"
       
    84   shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
       
    85 using a b
       
    86 proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
       
    87   case (Var z)
       
    88   have a1: "x \<noteq> y" by fact
       
    89   have a2: "atom x \<sharp> L" by fact
       
    90   show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS")
       
    91   proof -
       
    92     { -- {* Case 1.1 *}
       
    93       assume c1: "z = x"
       
    94       have "(1)": "?LHS = N[y::=L]" using c1 by simp
       
    95       have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp
       
    96       have "?LHS = ?RHS" using "(1)" "(2)" by simp
       
    97     }
       
    98     moreover 
       
    99     { -- {* Case 1.2 *}
       
   100       assume c2: "z = y" "z \<noteq> x" 
       
   101       
       
   102       have "?LHS = ?RHS" sorry
       
   103     }
       
   104     moreover 
       
   105     { -- {* Case 1.3 *}
       
   106       assume c3: "z \<noteq> x" "z \<noteq> y"
       
   107       
       
   108       have "?LHS = ?RHS" sorry
       
   109     }
       
   110     ultimately show "?LHS = ?RHS" by blast
       
   111   qed
       
   112 next
       
   113   case (Lam z M1) -- {* case 2: lambdas *}
       
   114   have ih: "\<lbrakk>x \<noteq> y; atom x \<sharp> L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact
       
   115   have a1: "x \<noteq> y" by fact
       
   116   have a2: "atom x \<sharp> L" by fact
       
   117   have fs: "atom z \<sharp> x" "atom z \<sharp> y" "atom z \<sharp> N" "atom z \<sharp> L" by fact+
       
   118   then have b: "atom z \<sharp> N[y::=L]" by (simp add: fresh_fact)
       
   119   show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") 
       
   120   proof - 
       
   121     have "?LHS = \<dots>" sorry
       
   122 
       
   123     also have "\<dots> = ?RHS" sorry
       
   124     finally show "?LHS = ?RHS" by simp
       
   125   qed
       
   126 next
       
   127   case (App M1 M2) -- {* case 3: applications *}
       
   128   then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp
       
   129 qed
       
   130 
       
   131 text {* 
       
   132   Again the strong induction principle enables Isabelle to find
       
   133   the proof of the substitution lemma automatically. 
       
   134 *}
       
   135 
       
   136 lemma substitution_lemma_version:  
       
   137   assumes asm: "x \<noteq> y" "atom x \<sharp> L"
       
   138   shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
       
   139   using asm 
       
   140 by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
       
   141    (auto simp add: fresh_fact forget)
       
   142 
       
   143 
       
   144 end