# HG changeset patch # User Cezary Kaliszyk # Date 1265797669 -3600 # Node ID 3acc7d25627acd5f3fb877f1885de836aa9d1d19 # Parent 9f6c606d5b594300e519bd0e17d0a9042b922e44 Some cleaning of proofs. diff -r 9f6c606d5b59 -r 3acc7d25627a Quot/QuotBase.thy --- a/Quot/QuotBase.thy Wed Feb 10 11:09:30 2010 +0100 +++ b/Quot/QuotBase.thy Wed Feb 10 11:27:49 2010 +0100 @@ -47,7 +47,7 @@ lemma eq_imp_rel: shows "equivp R \ a = b \ R a b" -by (simp add: equivp_reflp) + by (simp add: equivp_reflp) lemma identity_equivp: shows "equivp (op =)" @@ -252,15 +252,15 @@ assumes q: "Quotient R Abs Rep" and a: "R x1 x2" shows "R x1 (Rep (Abs x2))" -using a -by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) + using a + by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) lemma rep_abs_rsp_left: assumes q: "Quotient R Abs Rep" and a: "R x1 x2" shows "R (Rep (Abs x1)) x2" -using a -by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) + using a + by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) (* In the following theorem R1 can be instantiated with anything, but we know some of the types of the Rep and Abs functions; @@ -436,14 +436,14 @@ assumes a: "(R ===> (op =)) f g" shows "Ex1 (\x. x \ Respects R \ f x) = Ex1 (\x. x \ Respects R \ g x)" using a -by (simp add: Ex1_def in_respects) auto + by (simp add: Ex1_def in_respects) auto -(* 3 lemmas needed for cleaning of quantifiers *) +(* 2 lemmas needed for cleaning of quantifiers *) lemma all_prs: assumes a: "Quotient R absf repf" shows "Ball (Respects R) ((absf ---> id) f) = All f" using a unfolding Quotient_def Ball_def in_respects fun_map_def id_apply -by metis + by metis lemma ex_prs: assumes a: "Quotient R absf repf" @@ -458,45 +458,55 @@ where "Bex1_rel R P \ (\x \ Respects R. P x) \ (\x \ Respects R. \y \ Respects R. ((P x \ P y) \ (R x y)))" -(* TODO/FIXME: simplify the repetitions in this proof *) -lemma bexeq_rsp: +lemma bex1_rel_aux: + "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R x\ \ Bex1_rel R y" + unfolding Bex1_rel_def + apply (erule conjE)+ + apply (erule bexE) + apply rule + apply (rule_tac x="xa" in bexI) + apply metis + apply metis + apply rule+ + apply (erule_tac x="xaa" in ballE) + prefer 2 + apply (metis) + apply (erule_tac x="ya" in ballE) + prefer 2 + apply (metis) + apply (metis in_respects) + done + +lemma bex1_rel_aux2: + "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R y\ \ Bex1_rel R x" + unfolding Bex1_rel_def + apply (erule conjE)+ + apply (erule bexE) + apply rule + apply (rule_tac x="xa" in bexI) + apply metis + apply metis + apply rule+ + apply (erule_tac x="xaa" in ballE) + prefer 2 + apply (metis) + apply (erule_tac x="ya" in ballE) + prefer 2 + apply (metis) + apply (metis in_respects) + done + +lemma bex1_rel_rsp: assumes a: "Quotient R absf repf" shows "((R ===> op =) ===> op =) (Bex1_rel R) (Bex1_rel R)" -apply simp -unfolding Bex1_rel_def -apply rule -apply rule -apply rule -apply rule -apply (erule conjE)+ -apply (erule bexE) -apply rule -apply (rule_tac x="xa" in bexI) -apply metis -apply metis -apply rule+ -apply (erule_tac x="xb" in ballE) -prefer 2 -apply (metis) -apply (erule_tac x="ya" in ballE) -prefer 2 -apply (metis) -apply (metis in_respects) -apply (erule conjE)+ -apply (erule bexE) -apply rule -apply (rule_tac x="xa" in bexI) -apply metis -apply metis -apply rule+ -apply (erule_tac x="xb" in ballE) -prefer 2 -apply (metis) -apply (erule_tac x="ya" in ballE) -prefer 2 -apply (metis) -apply (metis in_respects) -done + apply simp + apply clarify + apply rule + apply (simp_all add: bex1_rel_aux bex1_rel_aux2) + apply (erule bex1_rel_aux2) + apply assumption + done + lemma ex1_prs: assumes a: "Quotient R absf repf" @@ -512,7 +522,6 @@ apply (erule conjE) apply (subgoal_tac "\y. R y y \ f (absf y) \ R x y") apply (rule_tac x="absf x" in exI) - apply (thin_tac "\x\Respects R. \y\Respects R. f (absf x) \ f (absf y) \ R x y") apply (simp) apply rule+ using a unfolding Quotient_def @@ -582,14 +591,14 @@ lemma if_prs: assumes q: "Quotient R Abs Rep" shows "Abs (If a (Rep b) (Rep c)) = If a b c" -using Quotient_abs_rep[OF q] by auto + using Quotient_abs_rep[OF q] by auto (* q not used *) lemma if_rsp: assumes q: "Quotient R Abs Rep" and a: "a1 = a2" "R b1 b2" "R c1 c2" shows "R (If a1 b1 c1) (If a2 b2 c2)" -using a by auto + using a by auto lemma let_prs: assumes q1: "Quotient R1 Abs1 Rep1" @@ -605,15 +614,13 @@ using apply_rsp[OF q1 a1] a2 by auto -(******************************************) -(* REST OF THE FILE IS UNUSED (until now) *) -(******************************************) +(*** REST OF THE FILE IS UNUSED (until now) ***) text {* lemma in_fun: shows "x \ ((f ---> g) s) = g (f x \ s)" by (simp add: mem_def) - + lemma respects_thm: shows "Respects (R1 ===> R2) f = (\x y. R1 x y \ R2 (f x) (f y))" unfolding Respects_def @@ -657,7 +664,7 @@ assumes q1: "Quotient R1 Abs1 Rep1" and q2: "Quotient R2 Abs2 Rep2" and r1: "Respects (R1 ===> R2) f" - and r2: "Respects (R1 ===> R2) g" + and r2: "Respects (R1 ===> R2) g" shows "((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g) = (\x y. R1 x y \ R2 (f x) (g y))" apply(rule_tac iffI) apply(rule)+ @@ -689,11 +696,6 @@ shows "(R1 ===> R2) f1 f2 \ (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))" using r1 r2 by auto -(* ask peter what are literal_case *) -(* literal_case_PRS *) -(* literal_case_RSP *) -(* Cez: !f x. literal_case f x = f x *) - (* We use id_simps which includes id_apply; so these 2 theorems can be removed *) lemma id_prs: assumes q: "Quotient R Abs Rep" diff -r 9f6c606d5b59 -r 3acc7d25627a Quot/QuotMain.thy --- a/Quot/QuotMain.thy Wed Feb 10 11:09:30 2010 +0100 +++ b/Quot/QuotMain.thy Wed Feb 10 11:27:49 2010 +0100 @@ -70,28 +70,28 @@ theorem thm11: shows "R r r' = (abs r = abs r')" -unfolding abs_def -by (simp only: equivp[simplified equivp_def] lem7) + unfolding abs_def + by (simp only: equivp[simplified equivp_def] lem7) lemma rep_refl: shows "R (rep a) (rep a)" -unfolding rep_def -by (simp add: equivp[simplified equivp_def]) + unfolding rep_def + by (simp add: equivp[simplified equivp_def]) lemma rep_abs_rsp: shows "R f (rep (abs g)) = R f g" and "R (rep (abs g)) f = R g f" -by (simp_all add: thm10 thm11) + by (simp_all add: thm10 thm11) lemma Quotient: shows "Quotient R abs rep" -unfolding Quotient_def -apply(simp add: thm10) -apply(simp add: rep_refl) -apply(subst thm11[symmetric]) -apply(simp add: equivp[simplified equivp_def]) -done + unfolding Quotient_def + apply(simp add: thm10) + apply(simp add: rep_refl) + apply(subst thm11[symmetric]) + apply(simp add: equivp[simplified equivp_def]) + done end @@ -142,10 +142,10 @@ and QT_ex1: "Quot_True (Ex1 P) \ Quot_True P" and QT_lam: "Quot_True (\x. P x) \ (\x. Quot_True (P x))" and QT_ext: "(\x. Quot_True (a x) \ f x = g x) \ (Quot_True a \ f = g)" -by (simp_all add: Quot_True_def ext) + by (simp_all add: Quot_True_def ext) lemma QT_imp: "Quot_True a \ Quot_True b" -by (simp add: Quot_True_def) + by (simp add: Quot_True_def) text {* Tactics for proving the lifted theorems *} use "quotient_tacs.ML"