# HG changeset patch # User Christian Urban # Date 1257151169 -3600 # Node ID 264c7b9d19f4a3784db7dc4c32ae9a0d199e8c2f # Parent 77ff9624cfd6494d5d2a4ce89cf57047faa8412b# Parent e169a99c6ada483ec30c794e90c0906109514f97 merged diff -r 77ff9624cfd6 -r 264c7b9d19f4 FSet.thy --- a/FSet.thy Mon Nov 02 09:33:48 2009 +0100 +++ b/FSet.thy Mon Nov 02 09:39:29 2009 +0100 @@ -359,18 +359,10 @@ ML {* val abs = findabs rty (prop_of (atomize_thm @{thm list.induct})) *} ML {* val aps = findaps rty (prop_of (atomize_thm @{thm list.induct})) *} - -ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *} -ML {* val ind_r_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms ind_r_t *} -lemma app_prs_for_induct: "(ABS_fset ---> id) f (REP_fset T1) = f T1" - apply (simp add: fun_map.simps QUOT_TYPE_I_fset.thm10) -done - -ML {* val ind_r_l1 = eqsubst_thm @{context} @{thms app_prs_for_induct} ind_r_l *} -ML {* val ind_r_l2 = eqsubst_thm @{context} @{thms app_prs_for_induct} ind_r_l1 *} -ML {* val ind_r_l3 = eqsubst_thm @{context} @{thms app_prs_for_induct} ind_r_l2 *} -ML {* val ind_r_l4 = eqsubst_thm @{context} @{thms app_prs_for_induct} ind_r_l3 *} -ML {* val ind_r_a = simp_allex_prs @{context} quot ind_r_l4 *} +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 ind_r_l = repeat_eqsubst_thm @{context} (simp_app_prs_thms @ simp_lam_prs_thms) ind_r_t *} +ML {* val ind_r_a = simp_allex_prs @{context} quot ind_r_l *} ML {* val thm = @{thm FORALL_PRS[OF FUN_QUOTIENT[OF QUOTIENT_fset IDENTITY_QUOTIENT], symmetric]} *} ML {* val ind_r_a1 = eqsubst_thm @{context} [thm] ind_r_a *} ML {* val defs_sym = add_lower_defs @{context} defs *} 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 \ 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 \ 'b) \ ('a noption) \ ('b noption)" diff -r 77ff9624cfd6 -r 264c7b9d19f4 QuotMain.thy --- a/QuotMain.thy Mon Nov 02 09:33:48 2009 +0100 +++ b/QuotMain.thy Mon Nov 02 09:39:29 2009 +0100 @@ -506,6 +506,10 @@ else Const (@{const_name "Ex"}, ty) $ apply_subt (my_reg lthy rel rty) t end + | Const (@{const_name "op ="}, ty) $ t => + if needs_lift rty (fastype_of t) then + (tyRel (fastype_of t) rty rel lthy) $ t + else Const (@{const_name "op ="}, ty) $ (my_reg lthy rel rty t) | t1 $ t2 => (my_reg lthy rel rty t1) $ (my_reg lthy rel rty t2) | _ => trm *} @@ -518,16 +522,32 @@ (my_reg lthy rel rty (prop_of thm))) *} +lemma universal_twice: "(\x. (P x \ Q x)) \ ((\x. P x) \ (\x. Q x))" +apply (auto) +done + +lemma implication_twice: "(c \ a) \ (a \ b \ d) \ (a \ b) \ (c \ d)" +apply (auto) +done + ML {* -fun regularize thm rty rel rel_eqv lthy = +fun regularize thm rty rel rel_eqv rel_refl lthy = let val g = build_regularize_goal thm rty rel lthy; fun tac ctxt = - (asm_full_simp_tac ((Simplifier.context ctxt HOL_ss) addsimps + (ObjectLogic.full_atomize_tac) THEN' + REPEAT_ALL_NEW (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])])) THEN_ALL_NEW - (((rtac @{thm RIGHT_RES_FORALL_REGULAR}) THEN' (RANGE [fn _ => all_tac, atac]) THEN' - (MetisTools.metis_tac ctxt [])) ORELSE' (MetisTools.metis_tac ctxt [])); + (@{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}) + ]); val cthm = Goal.prove lthy [] [] g (fn x => tac (#context x) 1); in cthm OF [thm] @@ -551,14 +571,10 @@ ) *} -ML {* make_const_def *} - ML {* fun old_get_fun flag rty qty lthy ty = get_fun flag [(qty, rty)] lthy ty -*} -ML {* fun old_make_const_def nconst_bname otrm mx rty qty lthy = make_const_def nconst_bname otrm mx [(qty, rty)] lthy *} @@ -838,26 +854,25 @@ Abs(_, T, b) => findaps_all rty (subst_bound ((Free ("x", T)), b)) | (f $ a) => (findaps_all rty f @ findaps_all rty a) - | Free (_, (T as (Type (_, (_ :: _))))) => (if needs_lift rty T then [T] else []) + | Free (_, (T as (Type ("fun", (_ :: _))))) => (if needs_lift rty T then [T] else []) | _ => []; fun findaps rty tm = distinct (op =) (findaps_all rty tm) *} ML {* -fun make_simp_lam_prs_thm lthy quot_thm typ = +fun make_simp_prs_thm lthy quot_thm thm typ = let val (_, [lty, rty]) = dest_Type typ; val thy = ProofContext.theory_of lthy; val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty) val inst = [SOME lcty, NONE, SOME rcty]; - val lpi = Drule.instantiate' inst [] @{thm LAMBDA_PRS}; + val lpi = Drule.instantiate' inst [] thm; val tac = - (compose_tac (false, @{thm LAMBDA_PRS}, 2)) THEN_ALL_NEW + (compose_tac (false, lpi, 2)) THEN_ALL_NEW (quotient_tac quot_thm); val t = Goal.prove lthy [] [] (concl_of lpi) (fn _ => tac 1); - val ts = @{thm HOL.sym} OF [t] in - MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] ts + MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] t end *} @@ -916,11 +931,13 @@ val (trans2, reps_same, quot) = lookup_quot_thms lthy qty_name; val consts = lookup_quot_consts defs; val t_a = atomize_thm t; - val t_r = regularize t_a rty rel rel_eqv lthy; + val t_r = regularize t_a rty rel rel_eqv rel_refl lthy; val t_t = repabs lthy t_r consts rty qty quot rel_refl trans2 rsp_thms; val abs = findabs rty (prop_of t_a); - val simp_lam_prs_thms = map (make_simp_lam_prs_thm lthy quot) abs; - val t_l = repeat_eqsubst_thm lthy simp_lam_prs_thms t_t; + val aps = findaps rty (prop_of t_a); + val app_prs_thms = map (make_simp_prs_thm lthy quot @{thm APP_PRS}) aps; + val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs; + val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_t; val t_a = simp_allex_prs lthy quot t_l; val defs_sym = add_lower_defs lthy defs; val t_d = repeat_eqsubst_thm lthy defs_sym t_a; diff -r 77ff9624cfd6 -r 264c7b9d19f4 QuotScript.thy --- a/QuotScript.thy Mon Nov 02 09:33:48 2009 +0100 +++ b/QuotScript.thy Mon Nov 02 09:39:29 2009 +0100 @@ -274,7 +274,7 @@ lemma LAMBDA_PRS: assumes q1: "QUOTIENT R1 Abs1 Rep1" and q2: "QUOTIENT R2 Abs2 Rep2" - shows "(\x. f x) = (Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x)))" + shows "(Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) = (\x. f x)" unfolding expand_fun_eq using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2] by simp @@ -284,7 +284,16 @@ and q2: "QUOTIENT R2 Abs2 Rep2" shows "(\x. f x) = (Rep1 ---> Abs2) (\x. (Abs1 ---> Rep2) f x)" unfolding expand_fun_eq -by (subst LAMBDA_PRS[OF q1 q2]) (simp) +using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2] +by (simp) + +lemma APP_PRS: + assumes q1: "QUOTIENT R1 abs1 rep1" + and q2: "QUOTIENT R2 abs2 rep2" + shows "abs2 ((abs1 ---> rep2) f (rep1 x)) = f x" +unfolding expand_fun_eq +using QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2] +by simp (* Ask Peter: assumption q1 and q2 not used and lemma is the 'identity' *) lemma LAMBDA_RSP: