diff -r 8948319b190e -r 3e405c2fbe90 Quot/Examples/LamEx.thy --- a/Quot/Examples/LamEx.thy Wed Feb 10 11:31:43 2010 +0100 +++ b/Quot/Examples/LamEx.thy Wed Feb 10 11:31:53 2010 +0100 @@ -72,8 +72,6 @@ \ rLam a t \ rLam b s" - - (* should be automatic with new version of eqvt-machinery *) lemma alpha_eqvt: fixes pi::"name prm" @@ -253,12 +251,18 @@ \ P lam" by (lifting rlam.induct) +ML {* show_all_types := true *} + lemma perm_lam [simp]: fixes pi::"'a prm" shows "pi \ Var a = Var (pi \ a)" and "pi \ App t1 t2 = App (pi \ t1) (pi \ t2)" and "pi \ Lam a t = Lam (pi \ a) (pi \ t)" apply(lifting perm_rlam.simps) +ML_prf {* + List.last (map (symmetric o #def) (Quotient_Info.qconsts_dest @{context})); + List.last (map (Thm.varifyT o symmetric o #def) (Quotient_Info.qconsts_dest @{context})) +*} done instance lam::pt_name