# HG changeset patch # User haftmann # Date 1273660312 -7200 # Node ID 3763c349c8c12345c4d1a1bb5907333b7cf1f5f9 # Parent b897bd9ca71bfb1cb58069105e9c312b968bd2ac grouped local statements diff -r b897bd9ca71b -r 3763c349c8c1 src/HOL/Semiring_Normalization.thy --- a/src/HOL/Semiring_Normalization.thy Wed May 12 12:31:51 2010 +0200 +++ b/src/HOL/Semiring_Normalization.thy Wed May 12 12:31:52 2010 +0200 @@ -50,13 +50,18 @@ then show "w * y + x * z = w * z + x * y \ w = x \ y = z" by auto qed +text {* semiring normalization proper *} + setup Semiring_Normalizer.setup -lemma (in comm_semiring_1) semiring_ops: +context comm_semiring_1 +begin + +lemma semiring_ops: shows "TERM (x + y)" and "TERM (x * y)" and "TERM (x ^ n)" and "TERM 0" and "TERM 1" . -lemma (in comm_semiring_1) semiring_rules: +lemma semiring_rules: "(a * m) + (b * m) = (a + b) * m" "(a * m) + m = (a + 1) * m" "m + (a * m) = (a + 1) * m" @@ -96,32 +101,42 @@ "x ^ (Suc (2*n)) = x * ((x ^ n) * (x ^ n))" by (simp_all add: algebra_simps power_add power2_eq_square power_mult_distrib power_mult) -lemmas (in comm_semiring_1) normalizing_comm_semiring_1_axioms = +lemmas normalizing_comm_semiring_1_axioms = comm_semiring_1_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules] -declaration (in comm_semiring_1) +declaration {* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_semiring_1_axioms} *} -lemma (in comm_ring_1) ring_ops: shows "TERM (x- y)" and "TERM (- x)" . +end -lemma (in comm_ring_1) ring_rules: +context comm_ring_1 +begin + +lemma ring_ops: shows "TERM (x- y)" and "TERM (- x)" . + +lemma ring_rules: "- x = (- 1) * x" "x - y = x + (- y)" by (simp_all add: diff_minus) -lemmas (in comm_ring_1) normalizing_comm_ring_1_axioms = +lemmas normalizing_comm_ring_1_axioms = comm_ring_1_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules ring ops: ring_ops ring rules: ring_rules] -declaration (in comm_ring_1) +declaration {* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_ring_1_axioms} *} -lemma (in comm_semiring_1_cancel_norm) noteq_reduce: +end + +context comm_semiring_1_cancel_norm +begin + +lemma noteq_reduce: "a \ b \ c \ d \ (a * c) + (b * d) \ (a * d) + (b * c)" proof- have "a \ b \ c \ d \ \ (a = b \ c = d)" by simp @@ -131,7 +146,7 @@ by simp qed -lemma (in comm_semiring_1_cancel_norm) add_scale_eq_noteq: +lemma add_scale_eq_noteq: "\r \ 0 ; a = b \ c \ d\ \ a + (r * c) \ b + (r * d)" proof(clarify) assume nz: "r\ 0" and cnd: "c\d" @@ -141,29 +156,34 @@ thus "False" using add_mult_solve[of 0 d] nz cnd by simp qed -lemma (in comm_semiring_1_cancel_norm) add_0_iff: +lemma add_0_iff: "x = x + a \ a = 0" proof- have "a = 0 \ x + a = x + 0" using add_imp_eq[of x a 0] by auto thus "x = x + a \ a = 0" by (auto simp add: add_commute) qed -declare (in comm_semiring_1_cancel_norm) +declare normalizing_comm_semiring_1_axioms [normalizer del] -lemmas (in comm_semiring_1_cancel_norm) +lemmas normalizing_comm_semiring_1_cancel_norm_axioms = comm_semiring_1_cancel_norm_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules idom rules: noteq_reduce add_scale_eq_noteq] -declaration (in comm_semiring_1_cancel_norm) +declaration {* Semiring_Normalizer.semiring_funs @{thm normalizing_comm_semiring_1_cancel_norm_axioms} *} -declare (in idom) normalizing_comm_ring_1_axioms [normalizer del] +end -lemmas (in idom) normalizing_idom_axioms = idom_axioms [normalizer +context idom +begin + +declare normalizing_comm_ring_1_axioms [normalizer del] + +lemmas normalizing_idom_axioms = idom_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules ring ops: ring_ops @@ -171,15 +191,20 @@ idom rules: noteq_reduce add_scale_eq_noteq ideal rules: right_minus_eq add_0_iff] -declaration (in idom) +declaration {* Semiring_Normalizer.semiring_funs @{thm normalizing_idom_axioms} *} -lemma (in field) field_ops: +end + +context field +begin + +lemma field_ops: shows "TERM (x / y)" and "TERM (inverse x)" . -lemmas (in field) field_rules = divide_inverse inverse_eq_divide +lemmas field_rules = divide_inverse inverse_eq_divide -lemmas (in field) normalizing_field_axioms = +lemmas normalizing_field_axioms = field_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules @@ -190,16 +215,18 @@ idom rules: noteq_reduce add_scale_eq_noteq ideal rules: right_minus_eq add_0_iff] -declaration (in field) +declaration {* Semiring_Normalizer.field_funs @{thm normalizing_field_axioms} *} +end + hide_fact (open) normalizing_comm_semiring_1_axioms normalizing_comm_semiring_1_cancel_norm_axioms semiring_ops semiring_rules hide_fact (open) normalizing_comm_ring_1_axioms normalizing_idom_axioms ring_ops ring_rules -hide_fact (open) normalizing_field_axioms field_ops field_rules +hide_fact (open) normalizing_field_axioms field_ops field_rules hide_fact (open) add_scale_eq_noteq noteq_reduce