| 25722 |      1 | (* $Id$ *)
 | 
|  |      2 | 
 | 
| 25727 |      3 | theory VC_Condition
 | 
| 25722 |      4 | imports "../Nominal" 
 | 
|  |      5 | begin
 | 
|  |      6 | 
 | 
|  |      7 | text {* 
 | 
| 26262 |      8 |   We give here two examples showing that if we use the variable  
 | 
| 25722 |      9 |   convention carelessly in rule inductions, we might end 
 | 
|  |     10 |   up with faulty lemmas. The point is that the examples
 | 
|  |     11 |   are not variable-convention compatible and therefore in the 
 | 
| 25751 |     12 |   nominal data package one is protected from such bogus reasoning.
 | 
| 25722 |     13 | *}
 | 
|  |     14 | 
 | 
| 25751 |     15 | text {* We define alpha-equated lambda-terms as usual. *}
 | 
| 25722 |     16 | atom_decl name 
 | 
|  |     17 | 
 | 
|  |     18 | nominal_datatype lam = 
 | 
|  |     19 |     Var "name"
 | 
|  |     20 |   | App "lam" "lam"
 | 
|  |     21 |   | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
 | 
|  |     22 | 
 | 
|  |     23 | text {*
 | 
|  |     24 |   The inductive relation 'unbind' unbinds the top-most  
 | 
|  |     25 |   binders of a lambda-term; this relation is obviously  
 | 
|  |     26 |   not a function, since it does not respect alpha-      
 | 
|  |     27 |   equivalence. However as a relation 'unbind' is ok and     
 | 
|  |     28 |   a similar relation has been used in our formalisation 
 | 
|  |     29 |   of the algorithm W. *}
 | 
| 26055 |     30 | 
 | 
| 25722 |     31 | inductive
 | 
|  |     32 |   unbind :: "lam \<Rightarrow> name list \<Rightarrow> lam \<Rightarrow> bool" ("_ \<mapsto> _,_" [60,60,60] 60)
 | 
|  |     33 | where
 | 
|  |     34 |   u_var: "(Var a) \<mapsto> [],(Var a)"
 | 
|  |     35 | | u_app: "(App t1 t2) \<mapsto> [],(App t1 t2)"
 | 
|  |     36 | | u_lam: "t\<mapsto>xs,t' \<Longrightarrow> (Lam [x].t) \<mapsto> (x#xs),t'"
 | 
|  |     37 | 
 | 
|  |     38 | text {*
 | 
|  |     39 |   We can show that Lam [x].Lam [x].Var x unbinds to [x,x],Var x 
 | 
|  |     40 |   and also to [z,y],Var y (though the proof for the second is a 
 | 
|  |     41 |   bit clumsy). *} 
 | 
| 26055 |     42 | 
 | 
| 25722 |     43 | lemma unbind_lambda_lambda1: 
 | 
|  |     44 |   shows "Lam [x].Lam [x].(Var x)\<mapsto>[x,x],(Var x)"
 | 
|  |     45 | by (auto intro: unbind.intros)
 | 
|  |     46 | 
 | 
|  |     47 | lemma unbind_lambda_lambda2: 
 | 
|  |     48 |   shows "Lam [x].Lam [x].(Var x)\<mapsto>[y,z],(Var z)"
 | 
|  |     49 | proof -
 | 
|  |     50 |   have "Lam [x].Lam [x].(Var x) = Lam [y].Lam [z].(Var z)" 
 | 
|  |     51 |     by (auto simp add: lam.inject alpha calc_atm abs_fresh fresh_atm)
 | 
|  |     52 |   moreover
 | 
|  |     53 |   have "Lam [y].Lam [z].(Var z) \<mapsto> [y,z],(Var z)"
 | 
|  |     54 |     by (auto intro: unbind.intros)
 | 
|  |     55 |   ultimately 
 | 
|  |     56 |   show "Lam [x].Lam [x].(Var x)\<mapsto>[y,z],(Var z)" by simp
 | 
|  |     57 | qed
 | 
|  |     58 | 
 | 
|  |     59 | text {* Unbind is equivariant ...*}
 | 
|  |     60 | equivariance unbind
 | 
|  |     61 | 
 | 
|  |     62 | text {*
 | 
|  |     63 |   ... but it is not variable-convention compatible (see Urban, 
 | 
| 25751 |     64 |   Berghofer, Norrish [2007]). This condition requires for rule u_lam to 
 | 
|  |     65 |   have the binder x not being a free variable in this rule's conclusion. 
 | 
|  |     66 |   Because this condition is not satisfied, Isabelle will not derive a 
 | 
|  |     67 |   strong induction principle for 'unbind' - that means Isabelle does not 
 | 
|  |     68 |   allow us to use the variable convention in induction proofs over 'unbind'. 
 | 
|  |     69 |   We can, however, force Isabelle to derive the strengthening induction 
 | 
|  |     70 |   principle and see what happens. *}
 | 
| 25722 |     71 | 
 | 
|  |     72 | nominal_inductive unbind
 | 
|  |     73 |   sorry
 | 
|  |     74 | 
 | 
|  |     75 | text {*
 | 
|  |     76 |   To obtain a faulty lemma, we introduce the function 'bind'
 | 
| 25751 |     77 |   which takes a list of names and abstracts them away in 
 | 
| 25722 |     78 |   a given lambda-term. *}
 | 
| 26055 |     79 | 
 | 
| 25722 |     80 | fun 
 | 
|  |     81 |   bind :: "name list \<Rightarrow> lam \<Rightarrow> lam"
 | 
|  |     82 | where
 | 
|  |     83 |   "bind [] t = t"
 | 
|  |     84 | | "bind (x#xs) t = Lam [x].(bind xs t)"
 | 
|  |     85 | 
 | 
|  |     86 | text {*
 | 
|  |     87 |   Although not necessary for our main argument below, we can 
 | 
|  |     88 |   easily prove that bind undoes the unbinding. *}
 | 
| 26055 |     89 | 
 | 
| 25722 |     90 | lemma bind_unbind:
 | 
|  |     91 |   assumes a: "t \<mapsto> xs,t'"
 | 
|  |     92 |   shows "t = bind xs t'"
 | 
|  |     93 | using a by (induct) (auto)
 | 
|  |     94 | 
 | 
|  |     95 | text {*
 | 
|  |     96 |   The next lemma shows by induction on xs that if x is a free 
 | 
| 25867 |     97 |   variable in t, and x does not occur in xs, then x is a free 
 | 
| 25722 |     98 |   variable in bind xs t. In the nominal tradition we formulate    
 | 
|  |     99 |   'is a free variable in' as 'is not fresh for'. *}
 | 
| 26055 |    100 | 
 | 
| 25722 |    101 | lemma free_variable:
 | 
|  |    102 |   fixes x::"name"
 | 
|  |    103 |   assumes a: "\<not>(x\<sharp>t)" and b: "x\<sharp>xs"
 | 
|  |    104 |   shows "\<not>(x\<sharp>bind xs t)"
 | 
|  |    105 | using a b
 | 
|  |    106 | by (induct xs)
 | 
|  |    107 |    (auto simp add: fresh_list_cons abs_fresh fresh_atm)
 | 
|  |    108 | 
 | 
|  |    109 | text {*
 | 
| 25751 |    110 |   Now comes the first faulty lemma. It is derived using the     
 | 
| 25722 |    111 |   variable convention (i.e. our strong induction principle). 
 | 
| 26262 |    112 |   This faulty lemma states that if t unbinds to x#xs and t', 
 | 
| 25722 |    113 |   and x is a free variable in t', then it is also a free 
 | 
|  |    114 |   variable in bind xs t'. We show this lemma by assuming that 
 | 
|  |    115 |   the binder x is fresh w.r.t. to the xs unbound previously. *}   
 | 
|  |    116 | 
 | 
|  |    117 | lemma faulty1:
 | 
|  |    118 |   assumes a: "t\<mapsto>(x#xs),t'"
 | 
|  |    119 |   shows "\<not>(x\<sharp>t') \<Longrightarrow> \<not>(x\<sharp>bind xs t')"
 | 
|  |    120 | using a
 | 
|  |    121 | by (nominal_induct t xs'\<equiv>"x#xs" t' avoiding: xs rule: unbind.strong_induct)
 | 
|  |    122 |    (simp_all add: free_variable)
 | 
|  |    123 | 
 | 
|  |    124 | text {*
 | 
| 25751 |    125 |   Obviously this lemma is bogus. For example, in 
 | 
|  |    126 |   case Lam [x].Lam [x].(Var x) \<mapsto> [x,x],(Var x). 
 | 
| 26055 |    127 |   As a result, we can prove False. *}
 | 
|  |    128 | 
 | 
| 25722 |    129 | lemma false1:
 | 
|  |    130 |   shows "False"
 | 
|  |    131 | proof -
 | 
| 26932 |    132 |   fix x
 | 
| 25722 |    133 |   have "Lam [x].Lam [x].(Var x)\<mapsto>[x,x],(Var x)" 
 | 
|  |    134 |   and  "\<not>(x\<sharp>Var x)" by (simp_all add: unbind_lambda_lambda1 fresh_atm)
 | 
|  |    135 |   then have "\<not>(x\<sharp>(bind [x] (Var x)))" by (rule faulty1)
 | 
|  |    136 |   moreover 
 | 
|  |    137 |   have "x\<sharp>(bind [x] (Var x))" by (simp add: abs_fresh)
 | 
|  |    138 |   ultimately
 | 
|  |    139 |   show "False" by simp
 | 
|  |    140 | qed
 | 
|  |    141 |    
 | 
|  |    142 | text {* 
 | 
|  |    143 |   The next example is slightly simpler, but looks more
 | 
| 25751 |    144 |   contrived than 'unbind'. This example, called 'strip' just 
 | 
| 25722 |    145 |   strips off the top-most binders from lambdas. *}
 | 
|  |    146 | 
 | 
|  |    147 | inductive
 | 
|  |    148 |   strip :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<rightarrow> _" [60,60] 60)
 | 
|  |    149 | where
 | 
|  |    150 |   s_var: "(Var a) \<rightarrow> (Var a)"
 | 
|  |    151 | | s_app: "(App t1 t2) \<rightarrow> (App t1 t2)"
 | 
|  |    152 | | s_lam: "t \<rightarrow> t' \<Longrightarrow> (Lam [x].t) \<rightarrow> t'"
 | 
|  |    153 | 
 | 
|  |    154 | text {* 
 | 
|  |    155 |   The relation is equivariant but we have to use again 
 | 
| 26055 |    156 |   sorry to derive a strong induction principle. *}
 | 
|  |    157 | 
 | 
| 25722 |    158 | equivariance strip
 | 
|  |    159 | 
 | 
|  |    160 | nominal_inductive strip
 | 
|  |    161 |   sorry
 | 
|  |    162 | 
 | 
|  |    163 | text {*
 | 
| 25751 |    164 |   The second faulty lemma shows that a variable being fresh
 | 
| 26055 |    165 |   for a term is also fresh for this term after striping. *}
 | 
|  |    166 | 
 | 
| 25722 |    167 | lemma faulty2:
 | 
|  |    168 |   fixes x::"name"
 | 
|  |    169 |   assumes a: "t \<rightarrow> t'"
 | 
|  |    170 |   shows "x\<sharp>t \<Longrightarrow> x\<sharp>t'"
 | 
|  |    171 | using a
 | 
|  |    172 | by (nominal_induct t t'\<equiv>t' avoiding: t' rule: strip.strong_induct)
 | 
|  |    173 |    (auto simp add: abs_fresh)
 | 
|  |    174 | 
 | 
| 25751 |    175 | text {* Obviously Lam [x].Var x is a counter example to this lemma. *}
 | 
| 26055 |    176 | 
 | 
| 25722 |    177 | lemma false2:
 | 
|  |    178 |   shows "False"
 | 
|  |    179 | proof -
 | 
| 26932 |    180 |   fix x
 | 
| 25722 |    181 |   have "Lam [x].(Var x) \<rightarrow> (Var x)" by (auto intro: strip.intros)
 | 
|  |    182 |   moreover
 | 
|  |    183 |   have "x\<sharp>Lam [x].(Var x)" by (simp add: abs_fresh)
 | 
|  |    184 |   ultimately have "x\<sharp>(Var x)" by (simp only: faulty2)
 | 
|  |    185 |   then show "False" by (simp add: fresh_atm)
 | 
|  |    186 | qed 
 | 
| 26055 |    187 | 
 | 
| 26262 |    188 | text {* A similar effect can be achieved by using naive inversion 
 | 
|  |    189 |   on rules. *}
 | 
| 26095 |    190 | 
 | 
|  |    191 | lemma false3: 
 | 
|  |    192 |   shows "False"
 | 
|  |    193 | proof -
 | 
| 26932 |    194 |   fix x
 | 
| 26095 |    195 |   have "Lam [x].(Var x) \<rightarrow> (Var x)" by (auto intro: strip.intros)
 | 
| 26262 |    196 |   moreover obtain y::"name" where fc: "y\<sharp>x" by (rule exists_fresh) (rule fin_supp)
 | 
| 26095 |    197 |   then have "Lam [x].(Var x) = Lam [y].(Var y)"
 | 
|  |    198 |     by (simp add: lam.inject alpha calc_atm fresh_atm)
 | 
|  |    199 |   ultimately have "Lam [y].(Var y) \<rightarrow> (Var x)" by simp
 | 
| 26262 |    200 |   then have "Var y \<rightarrow> Var x" using fc
 | 
|  |    201 |     by (cases rule: strip.strong_cases[where x=y])
 | 
| 26095 |    202 |        (simp_all add: lam.inject alpha abs_fresh)
 | 
| 26262 |    203 |   then show "False" using fc
 | 
| 26095 |    204 |     by (cases) (simp_all add: lam.inject fresh_atm)
 | 
|  |    205 | qed
 | 
|  |    206 | 
 | 
| 26262 |    207 | text {*
 | 
|  |    208 |   Moral: Who would have thought that the variable convention 
 | 
|  |    209 |   is in general an unsound reasoning principle.
 | 
|  |    210 |   *}
 | 
|  |    211 | 
 | 
| 25722 |    212 | end |