diff -r ef362328b931 -r 12990a6dddcb src/HOL/Nominal/Examples/VC_Condition.thy --- a/src/HOL/Nominal/Examples/VC_Condition.thy Tue Oct 08 13:13:53 2024 +0200 +++ b/src/HOL/Nominal/Examples/VC_Condition.thy Tue Oct 08 15:02:17 2024 +0200 @@ -27,11 +27,11 @@ of the algorithm W.\ inductive - unbind :: "lam \ name list \ lam \ bool" (\_ \ _,_\ [60,60,60] 60) + unbind :: "lam \ name list \ lam \ bool" (\_ \ _,_\ [60,60,60] 60) where - u_var: "(Var a) \ [],(Var a)" -| u_app: "(App t1 t2) \ [],(App t1 t2)" -| u_lam: "t\xs,t' \ (Lam [x].t) \ (x#xs),t'" + u_var: "(Var a) \ [],(Var a)" +| u_app: "(App t1 t2) \ [],(App t1 t2)" +| u_lam: "t\xs,t' \ (Lam [x].t) \ (x#xs),t'" text \ We can show that Lam [x].Lam [x].Var x unbinds to [x,x],Var x @@ -39,19 +39,19 @@ bit clumsy).\ lemma unbind_lambda_lambda1: - shows "Lam [x].Lam [x].(Var x)\[x,x],(Var x)" + shows "Lam [x].Lam [x].(Var x)\[x,x],(Var x)" by (auto intro: unbind.intros) lemma unbind_lambda_lambda2: - shows "Lam [x].Lam [x].(Var x)\[y,z],(Var z)" + shows "Lam [x].Lam [x].(Var x)\[y,z],(Var z)" proof - have "Lam [x].Lam [x].(Var x) = Lam [y].Lam [z].(Var z)" by (auto simp add: lam.inject alpha calc_atm abs_fresh fresh_atm) moreover - have "Lam [y].Lam [z].(Var z) \ [y,z],(Var z)" + have "Lam [y].Lam [z].(Var z) \ [y,z],(Var z)" by (auto intro: unbind.intros) ultimately - show "Lam [x].Lam [x].(Var x)\[y,z],(Var z)" by simp + show "Lam [x].Lam [x].(Var x)\[y,z],(Var z)" by simp qed text \Unbind is equivariant ...\ @@ -86,7 +86,7 @@ easily prove that bind undoes the unbinding.\ lemma bind_unbind: - assumes a: "t \ xs,t'" + assumes a: "t \ xs,t'" shows "t = bind xs t'" using a by (induct) (auto) @@ -113,7 +113,7 @@ the binder x is fresh w.r.t. to the xs unbound previously.\ lemma faulty1: - assumes a: "t\(x#xs),t'" + assumes a: "t\(x#xs),t'" shows "\(x\t') \ \(x\bind xs t')" using a by (nominal_induct t xs'\"x#xs" t' avoiding: xs rule: unbind.strong_induct) @@ -121,14 +121,14 @@ text \ Obviously this lemma is bogus. For example, in - case Lam [x].Lam [x].(Var x) \ [x,x],(Var x). + case Lam [x].Lam [x].(Var x) \ [x,x],(Var x). As a result, we can prove False.\ lemma false1: shows "False" proof - fix x - have "Lam [x].Lam [x].(Var x)\[x,x],(Var x)" + have "Lam [x].Lam [x].(Var x)\[x,x],(Var x)" and "\(x\Var x)" by (simp_all add: unbind_lambda_lambda1 fresh_atm) then have "\(x\(bind [x] (Var x)))" by (rule faulty1) moreover