diff -r 77ff9624cfd6 -r 264c7b9d19f4 LamEx.thy --- a/LamEx.thy Mon Nov 02 09:33:48 2009 +0100 +++ b/LamEx.thy Mon Nov 02 09:39:29 2009 +0100 @@ -196,6 +196,7 @@ thm rlam.inject thm pi_var + lemma var_inject: shows "(Var a = Var b) = (a = b)" @@ -221,19 +222,54 @@ -ML {* val t_a = atomize_thm @{thm alpha.cases} *} +ML {* val t_a = atomize_thm @{thm alpha.induct} *} +(* prove {* build_regularize_goal t_a rty rel @{context} *} +ML_prf {* fun tac ctxt = + (FIRST' [ + rtac rel_refl, + atac, + rtac @{thm universal_twice}, + (rtac @{thm impI} THEN' atac), + rtac @{thm implication_twice}, + EqSubst.eqsubst_tac ctxt [0] + [(@{thm equiv_res_forall} OF [rel_eqv]), + (@{thm equiv_res_exists} OF [rel_eqv])], + (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl), + (rtac @{thm RIGHT_RES_FORALL_REGULAR}) + ]); + *} + apply (tactic {* tac @{context} 1 *}) *) ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *} ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *} -ML {* val abs = findabs rty (prop_of t_a); *} -ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *} -ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *} + +ML {* val abs = findabs rty (prop_of (atomize_thm @{thm alpha.induct})) *} +ML {* val aps = findaps rty (prop_of (atomize_thm @{thm alpha.induct})) *} +ML {* val aps = @{typ "LamEx.rlam \<Rightarrow> bool"} :: aps; *} +ML {* val simp_app_prs_thms = map (make_simp_prs_thm @{context} quot @{thm APP_PRS}) aps *} +ML {* val simp_lam_prs_thms = map (make_simp_prs_thm @{context} quot @{thm LAMBDA_PRS}) abs *} +ML {* val t_l = repeat_eqsubst_thm @{context} (simp_app_prs_thms @ simp_lam_prs_thms) t_t *} ML {* val t_a = simp_allex_prs @{context} quot t_l *} +ML {* val thm = + @{thm FORALL_PRS[OF FUN_QUOTIENT[OF QUOTIENT_lam FUN_QUOTIENT[OF QUOTIENT_lam IDENTITY_QUOTIENT]], symmetric]} *} +ML {* val t_a1 = eqsubst_thm @{context} [thm] t_a *} ML {* val defs_sym = add_lower_defs @{context} defs; *} -ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_a *} +ML {* val t_d = repeat_eqsubst_thm @{context} defs_sym t_a1 *} ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *} -ML {* ObjectLogic.rulify t_r *} +ML {* val t_r1 = repeat_eqsubst_thm @{context} @{thms fun_map.simps} t_r *} +ML {* val t_r2 = repeat_eqsubst_thm @{context} @{thms QUOT_TYPE_I_lam.thm10} t_r1 *} +ML {* val t_r3 = repeat_eqsubst_thm @{context} @{thms id_apply} t_r2 *} + +ML {* val alpha_induct = ObjectLogic.rulify t_r3 *} +local_setup {* + Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd +*} +lemma + assumes a: "(a::lam) = b" + shows "False" + using a + apply (rule alpha_induct) fun option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"