diff -r a609eea06119 -r c32bdb3f0a68 Nominal/Nominal2.thy --- a/Nominal/Nominal2.thy Thu Sep 22 11:42:55 2011 +0200 +++ b/Nominal/Nominal2.thy Sat Dec 17 16:36:51 2011 +0000 @@ -443,14 +443,19 @@ val qfv_supp_thms = prove_fv_supp qtys (flat qtrms) qfvs qfv_bns qalpha_bns qfv_defs qeq_iffs qperm_simps qfv_qbn_eqvts qinduct (flat raw_bclauses) lthyC + |> map Drule.eta_contraction_rule (* postprocessing of eq and fv theorems *) val qeq_iffs' = qeq_iffs |> map (simplify (HOL_basic_ss addsimps qfv_supp_thms)) |> map (simplify (HOL_basic_ss addsimps @{thms prod_fv_supp prod_alpha_eq Abs_eq_iff[symmetric]})) + (* filters the theormes that are of the form "qfv = supp" *) + fun is_qfv_thm (@{term Trueprop} $ (Const (@{const_name HOL.eq}, _) $ lhs $ _)) = member (op=) qfvs lhs + | is_qfv_thm _ = false + val qsupp_constrs = qfv_defs - |> map (simplify (HOL_basic_ss addsimps (take (length qfvs) qfv_supp_thms))) + |> map (simplify (HOL_basic_ss addsimps (filter (is_qfv_thm o prop_of) qfv_supp_thms))) val transform_thm = @{lemma "x = y \ a \ x \ a \ y" by simp} val transform_thms =