--- a/Nominal/Fv.thy Mon Mar 22 17:21:27 2010 +0100
+++ b/Nominal/Fv.thy Mon Mar 22 18:29:29 2010 +0100
@@ -679,19 +679,24 @@
*}
ML {*
-fun build_alpha_refl_gl fv_alphas_lst =
+fun build_alpha_refl_gl fv_alphas_lst alphas =
let
- val (fvs_alphas, ls) = split_list fv_alphas_lst;
+ val (fvs_alphas, _) = split_list fv_alphas_lst;
val (_, alpha_ts) = split_list fvs_alphas;
val tys = map (domain_type o fastype_of) alpha_ts;
val names = Datatype_Prop.make_tnames tys;
val args = map Free (names ~~ tys);
+ fun find_alphas ty x =
+ domain_type (fastype_of x) = ty;
fun mk_alpha_refl arg (_, alpha) = alpha $ arg $ arg;
- fun refl_eq_arg ((alpha, arg), l) =
- foldr1 HOLogic.mk_conj
- ((alpha $ arg $ arg) ::
- (map (mk_alpha_refl arg) l))
- val eqs = map refl_eq_arg ((alpha_ts ~~ args) ~~ ls)
+ fun refl_eq_arg (ty, arg) =
+ let
+ val rel_alphas = filter (find_alphas ty) alphas;
+ in
+ map (fn x => x $ arg $ arg) rel_alphas
+ end;
+ (* Flattening loses the induction structure *)
+ val eqs = map (foldr1 HOLogic.mk_conj) (map refl_eq_arg (tys ~~ args))
in
(names, HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj eqs))
end
@@ -708,9 +713,9 @@
*}
ML {*
-fun build_alpha_refl fv_alphas_lst induct eq_iff ctxt =
+fun build_alpha_refl fv_alphas_lst alphas induct eq_iff ctxt =
let
- val (names, gl) = build_alpha_refl_gl fv_alphas_lst;
+ val (names, gl) = build_alpha_refl_gl fv_alphas_lst alphas;
val refl_conj = Goal.prove ctxt names [] gl (fn _ => reflp_tac induct eq_iff ctxt 1);
in
HOLogic.conj_elims refl_conj
@@ -825,23 +830,25 @@
split_conjs THEN_ALL_NEW REPEAT o (eetac @{thm exi_sum} ctxt) THEN_ALL_NEW split_conjs
THEN_ALL_NEW (asm_full_simp_tac (HOL_ss addsimps (term_inj @ distinct)))
THEN_ALL_NEW split_conjs THEN_ALL_NEW
- TRY o (etac @{thm alpha_gen_compose_trans} THEN' RANGE[atac]) THEN_ALL_NEW
+ TRY o (etac @{thm alpha_gen_compose_trans2} ORELSE' etac @{thm alpha_gen_compose_trans}) THEN_ALL_NEW
(asm_full_simp_tac (HOL_ss addsimps (all_eqvts ctxt @ eqvt @ term_inj @ distinct)))
*}
-lemma transp_aux:
+lemma transpI:
"(\<And>xa ya. R xa ya \<longrightarrow> (\<forall>z. R ya z \<longrightarrow> R xa z)) \<Longrightarrow> transp R"
unfolding transp_def
by blast
ML {*
fun equivp_tac reflps symps transps =
+ let val _ = tracing (PolyML.makestring (reflps, symps, transps)) in
simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def})
THEN' rtac conjI THEN' rtac allI THEN'
resolve_tac reflps THEN'
rtac conjI THEN' rtac allI THEN' rtac allI THEN'
resolve_tac symps THEN'
- rtac @{thm transp_aux} THEN' resolve_tac transps
+ rtac @{thm transpI} THEN' resolve_tac transps
+ end
*}
ML {*