# HG changeset patch # User ballarin # Date 1266193626 -3600 # Node ID 8f1e60d9f7cc1eea752a3855a7d7c0199c2b4c45 # Parent 055934ed89b0c921915e4ef3ef65cbf06526556d Tuned interpretation proofs. diff -r 055934ed89b0 -r 8f1e60d9f7cc src/HOL/Algebra/UnivPoly.thy --- a/src/HOL/Algebra/UnivPoly.thy Thu Feb 11 21:00:36 2010 +0100 +++ b/src/HOL/Algebra/UnivPoly.thy Mon Feb 15 01:27:06 2010 +0100 @@ -1342,14 +1342,8 @@ text {* Interpretation of ring homomorphism lemmas. *} sublocale UP_univ_prop < ring_hom_cring P S Eval - apply (unfold Eval_def) - apply intro_locales - apply (rule ring_hom_cring.axioms) - apply (rule ring_hom_cring.intro) - apply unfold_locales - apply (rule eval_ring_hom) - apply rule - done + unfolding Eval_def + by unfold_locales (fast intro: eval_ring_hom) lemma (in UP_cring) monom_pow: assumes R: "a \ carrier R" @@ -1436,10 +1430,7 @@ lemma ring_homD: assumes Phi: "Phi \ ring_hom P S" shows "ring_hom_cring P S Phi" -proof (rule ring_hom_cring.intro) - show "ring_hom_cring_axioms P S Phi" - by (rule ring_hom_cring_axioms.intro) (rule Phi) -qed unfold_locales + by unfold_locales (rule Phi) theorem UP_universal_property: assumes S: "s \ carrier S" @@ -1759,9 +1750,9 @@ shows "eval R R id a (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) = \" (is "eval R R id a ?g = _") proof - - interpret UP_pre_univ_prop R R id proof qed simp + interpret UP_pre_univ_prop R R id by unfold_locales simp have eval_ring_hom: "eval R R id a \ ring_hom P R" using eval_ring_hom [OF a] by simp - interpret ring_hom_cring P R "eval R R id a" proof qed (simp add: eval_ring_hom) + interpret ring_hom_cring P R "eval R R id a" by unfold_locales (rule eval_ring_hom) have mon1_closed: "monom P \\<^bsub>R\<^esub> 1 \ carrier P" and mon0_closed: "monom P a 0 \ carrier P" and min_mon0_closed: "\\<^bsub>P\<^esub> monom P a 0 \ carrier P" diff -r 055934ed89b0 -r 8f1e60d9f7cc src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Thu Feb 11 21:00:36 2010 +0100 +++ b/src/HOL/Lattices.thy Mon Feb 15 01:27:06 2010 +0100 @@ -301,8 +301,7 @@ lemma dual_bounded_lattice: "bounded_lattice (op \) (op >) (op \) (op \) \ \" - by (rule bounded_lattice.intro, rule dual_lattice) - (unfold_locales, auto simp add: less_le_not_le) + by unfold_locales (auto simp add: less_le_not_le) lemma inf_bot_left [simp]: "\ \ x = \"