Quot/Examples/LamEx.thy
changeset 882 6a8858ba01f6
parent 881 2cc520457e37
child 883 99e811fc1366
--- a/Quot/Examples/LamEx.thy	Thu Jan 14 18:35:38 2010 +0100
+++ b/Quot/Examples/LamEx.thy	Thu Jan 14 18:41:50 2010 +0100
@@ -287,7 +287,10 @@
   next
     case (3 name lam pi a)
     have b: "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" by fact
-    obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)" sorry
+    obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
+      apply(rule exists_fresh[of "(a, pi\<bullet>name, pi\<bullet>lam)"])
+      apply(simp_all add: fs_name1)
+      done
     from b fr have p: "P a (Lam c (([(c, pi\<bullet>name)]@pi)\<bullet>lam))" 
       apply -
       apply(rule a3)