# HG changeset patch # User Cezary Kaliszyk # Date 1264519851 -3600 # Node ID 46cc6708c3b33557440e28b7657dcca46836791e # Parent de9e5faf1f03ff12d15ed90c5150df6d3ad490c6 Bex1_Bexeq_regular. diff -r de9e5faf1f03 -r 46cc6708c3b3 Quot/Examples/SigmaEx.thy --- a/Quot/Examples/SigmaEx.thy Tue Jan 26 15:59:04 2010 +0100 +++ b/Quot/Examples/SigmaEx.thy Tue Jan 26 16:30:51 2010 +0100 @@ -85,14 +85,14 @@ \ fpar\Respects (op = ===> op = ===> alpha_method ===> op =). \ fsgm\Respects (op = ===> (op = ===> alpha_obj) ===> op =). - Bexeq - (prod_rel (alpha_obj ===> op =) + Bex1 + (Respects (prod_rel (alpha_obj ===> op =) (prod_rel (list_rel (prod_rel (op =) alpha_method) ===> op =) (prod_rel ((prod_rel (op =) alpha_method) ===> op =) (alpha_method ===> op =) ) ) - ) + )) (\ (hom_o\robj \ 'a, hom_d\(char list \ rmethod) list \ 'b, hom_e\char list \ rmethod \ 'c, hom_m\rmethod \ 'd). @@ -128,6 +128,16 @@ sorry +lemma bex1_bex1reg: "(\!x\Respects R. P x) \ (Bexeq R (\x. P x))" +apply (simp add: Bex1_def Ex1_def Bexeq_def in_respects) +apply clarify +apply auto +apply (rule bexI) +apply assumption +apply (simp add: in_respects) +apply (simp add: in_respects) +apply auto +done lemma liftd: " Ex1 (\(hom_o, hom_d, hom_e, hom_m). diff -r de9e5faf1f03 -r 46cc6708c3b3 Quot/QuotBase.thy --- a/Quot/QuotBase.thy Tue Jan 26 15:59:04 2010 +0100 +++ b/Quot/QuotBase.thy Tue Jan 26 16:30:51 2010 +0100 @@ -535,6 +535,17 @@ apply metis done +lemma bex1_bexeq_reg: "(\!x\Respects R. P x) \ (Bexeq R (\x. P x))" + apply (simp add: Bex1_def Ex1_def Bexeq_def in_respects) + apply clarify + apply auto + apply (rule bexI) + apply assumption + apply (simp add: in_respects) + apply (simp add: in_respects) + apply auto + done + section {* Various respects and preserve lemmas *} lemma quot_rel_rsp: diff -r de9e5faf1f03 -r 46cc6708c3b3 Quot/quotient_tacs.ML --- a/Quot/quotient_tacs.ML Tue Jan 26 15:59:04 2010 +0100 +++ b/Quot/quotient_tacs.ML Tue Jan 26 16:30:51 2010 +0100 @@ -132,7 +132,7 @@ 0. preliminary simplification step according to ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range - 1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left) + 1. eliminating simple Ball/Bex/Bex1 instances (ball_reg_right bex_reg_left) 2. monos @@ -159,7 +159,7 @@ in simp_tac simpset THEN' REPEAT_ALL_NEW (CHANGED o FIRST' - [resolve_tac @{thms ball_reg_right bex_reg_left}, + [resolve_tac @{thms ball_reg_right bex_reg_left bex1_bexeq_reg}, resolve_tac (Inductive.get_monos ctxt), resolve_tac @{thms ball_all_comm bex_ex_comm}, resolve_tac eq_eqvs, diff -r de9e5faf1f03 -r 46cc6708c3b3 Quot/quotient_term.ML --- a/Quot/quotient_term.ML Tue Jan 26 15:59:04 2010 +0100 +++ b/Quot/quotient_term.ML Tue Jan 26 16:30:51 2010 +0100 @@ -476,6 +476,16 @@ val needrel = Syntax.check_term ctxt (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) in if resrel <> needrel + then term_mismatch "regularize(bex1_res)" ctxt resrel needrel + else mk_bexeq $ resrel $ subtrm + end + + | (Const (@{const_name "Bex1"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t, Const (@{const_name "Ex1"}, ty') $ t') => + let + val subtrm = apply_subt (regularize_trm ctxt) (t, t') + val needrel = Syntax.check_term ctxt (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) + in + if resrel <> needrel then term_mismatch "regularize(bex1)" ctxt resrel needrel else mk_bexeq $ resrel $ subtrm end