Theory HOL-Decision_Procs.Algebra_Aux
section ‹Things that can be added to the Algebra library›
theory Algebra_Aux
imports "HOL-Algebra.Ring"
begin
definition of_natural :: "('a, 'm) ring_scheme ⇒ nat ⇒ 'a" ("«_»⇩ℕı")
where "«n»⇩ℕ⇘R⇙ = ((⊕⇘R⇙) 𝟭⇘R⇙ ^^ n) 𝟬⇘R⇙"
definition of_integer :: "('a, 'm) ring_scheme ⇒ int ⇒ 'a" ("«_»ı")
where "«i»⇘R⇙ = (if 0 ≤ i then «nat i»⇩ℕ⇘R⇙ else ⊖⇘R⇙ «nat (- i)»⇩ℕ⇘R⇙)"
context ring
begin
lemma of_nat_0 [simp]: "«0»⇩ℕ = 𝟬"
by (simp add: of_natural_def)
lemma of_nat_Suc [simp]: "«Suc n»⇩ℕ = 𝟭 ⊕ «n»⇩ℕ"
by (simp add: of_natural_def)
lemma of_int_0 [simp]: "«0» = 𝟬"
by (simp add: of_integer_def)
lemma of_nat_closed [simp]: "«n»⇩ℕ ∈ carrier R"
by (induct n) simp_all
lemma of_int_closed [simp]: "«i» ∈ carrier R"
by (simp add: of_integer_def)
lemma of_int_minus [simp]: "«- i» = ⊖ «i»"
by (simp add: of_integer_def)
lemma of_nat_add [simp]: "«m + n»⇩ℕ = «m»⇩ℕ ⊕ «n»⇩ℕ"
by (induct m) (simp_all add: a_ac)
lemma of_nat_diff [simp]: "n ≤ m ⟹ «m - n»⇩ℕ = «m»⇩ℕ ⊖ «n»⇩ℕ"
proof (induct m arbitrary: n)
case 0
then show ?case by (simp add: minus_eq)
next
case Suc': (Suc m)
show ?case
proof (cases n)
case 0
then show ?thesis
by (simp add: minus_eq)
next
case (Suc k)
with Suc' have "«Suc m - Suc k»⇩ℕ = «m»⇩ℕ ⊖ «k»⇩ℕ" by simp
also have "… = 𝟭 ⊕ ⊖ 𝟭 ⊕ («m»⇩ℕ ⊖ «k»⇩ℕ)"
by (simp add: r_neg)
also have "… = «Suc m»⇩ℕ ⊖ «Suc k»⇩ℕ"
by (simp add: minus_eq minus_add a_ac)
finally show ?thesis using Suc by simp
qed
qed
lemma of_int_add [simp]: "«i + j» = «i» ⊕ «j»"
proof (cases "0 ≤ i")
case True
show ?thesis
proof (cases "0 ≤ j")
case True
with ‹0 ≤ i› show ?thesis
by (simp add: of_integer_def nat_add_distrib)
next
case False
show ?thesis
proof (cases "0 ≤ i + j")
case True
then have "«i + j» = «nat (i - (- j))»⇩ℕ"
by (simp add: of_integer_def)
also from True ‹0 ≤ i› ‹¬ 0 ≤ j›
have "nat (i - (- j)) = nat i - nat (- j)"
by (simp add: nat_diff_distrib)
finally show ?thesis using True ‹0 ≤ i› ‹¬ 0 ≤ j›
by (simp add: minus_eq of_integer_def)
next
case False
then have "«i + j» = ⊖ «nat (- j - i)»⇩ℕ"
by (simp add: of_integer_def) (simp only: diff_conv_add_uminus add_ac)
also from False ‹0 ≤ i› ‹¬ 0 ≤ j›
have "nat (- j - i) = nat (- j) - nat i"
by (simp add: nat_diff_distrib)
finally show ?thesis using False ‹0 ≤ i› ‹¬ 0 ≤ j›
by (simp add: minus_eq minus_add a_ac of_integer_def)
qed
qed
next
case False
show ?thesis
proof (cases "0 ≤ j")
case True
show ?thesis
proof (cases "0 ≤ i + j")
case True
then have "«i + j» = «nat (j - (- i))»⇩ℕ"
by (simp add: of_integer_def add_ac)
also from True ‹¬ 0 ≤ i› ‹0 ≤ j›
have "nat (j - (- i)) = nat j - nat (- i)"
by (simp add: nat_diff_distrib)
finally show ?thesis using True ‹¬ 0 ≤ i› ‹0 ≤ j›
by (simp add: minus_eq minus_add a_ac of_integer_def)
next
case False
then have "«i + j» = ⊖ «nat (- i - j)»⇩ℕ"
by (simp add: of_integer_def)
also from False ‹¬ 0 ≤ i› ‹0 ≤ j›
have "nat (- i - j) = nat (- i) - nat j"
by (simp add: nat_diff_distrib)
finally show ?thesis using False ‹¬ 0 ≤ i› ‹0 ≤ j›
by (simp add: minus_eq minus_add of_integer_def)
qed
next
case False
with ‹¬ 0 ≤ i› show ?thesis
by (simp add: of_integer_def nat_add_distrib minus_add diff_conv_add_uminus
del: add_uminus_conv_diff uminus_add_conv_diff)
qed
qed
lemma of_int_diff [simp]: "«i - j» = «i» ⊖ «j»"
by (simp only: diff_conv_add_uminus of_int_add) (simp add: minus_eq)
lemma of_nat_mult [simp]: "«i * j»⇩ℕ = «i»⇩ℕ ⊗ «j»⇩ℕ"
by (induct i) (simp_all add: l_distr)
lemma of_int_mult [simp]: "«i * j» = «i» ⊗ «j»"
proof (cases "0 ≤ i")
case True
show ?thesis
proof (cases "0 ≤ j")
case True
with ‹0 ≤ i› show ?thesis
by (simp add: of_integer_def nat_mult_distrib zero_le_mult_iff)
next
case False
with ‹0 ≤ i› show ?thesis
by (simp add: of_integer_def zero_le_mult_iff
minus_mult_right nat_mult_distrib r_minus
del: minus_mult_right [symmetric])
qed
next
case False
show ?thesis
proof (cases "0 ≤ j")
case True
with ‹¬ 0 ≤ i› show ?thesis
by (simp add: of_integer_def zero_le_mult_iff
minus_mult_left nat_mult_distrib l_minus
del: minus_mult_left [symmetric])
next
case False
with ‹¬ 0 ≤ i› show ?thesis
by (simp add: of_integer_def zero_le_mult_iff
minus_mult_minus [of i j, symmetric] nat_mult_distrib
l_minus r_minus
del: minus_mult_minus
minus_mult_left [symmetric] minus_mult_right [symmetric])
qed
qed
lemma of_int_1 [simp]: "«1» = 𝟭"
by (simp add: of_integer_def)
lemma of_int_2: "«2» = 𝟭 ⊕ 𝟭"
by (simp add: of_integer_def numeral_2_eq_2)
lemma minus_0_r [simp]: "x ∈ carrier R ⟹ x ⊖ 𝟬 = x"
by (simp add: minus_eq)
lemma minus_0_l [simp]: "x ∈ carrier R ⟹ 𝟬 ⊖ x = ⊖ x"
by (simp add: minus_eq)
lemma eq_diff0:
assumes "x ∈ carrier R" "y ∈ carrier R"
shows "x ⊖ y = 𝟬 ⟷ x = y"
proof
assume "x ⊖ y = 𝟬"
with assms have "x ⊕ (⊖ y ⊕ y) = y"
by (simp add: minus_eq a_assoc [symmetric])
with assms show "x = y" by (simp add: l_neg)
next
assume "x = y"
with assms show "x ⊖ y = 𝟬" by (simp add: minus_eq r_neg)
qed
lemma power2_eq_square: "x ∈ carrier R ⟹ x [^] (2::nat) = x ⊗ x"
by (simp add: numeral_eq_Suc)
lemma eq_neg_iff_add_eq_0:
assumes "x ∈ carrier R" "y ∈ carrier R"
shows "x = ⊖ y ⟷ x ⊕ y = 𝟬"
proof
assume "x = ⊖ y"
with assms show "x ⊕ y = 𝟬" by (simp add: l_neg)
next
assume "x ⊕ y = 𝟬"
with assms have "x ⊕ (y ⊕ ⊖ y) = 𝟬 ⊕ ⊖ y"
by (simp add: a_assoc [symmetric])
with assms show "x = ⊖ y"
by (simp add: r_neg)
qed
lemma neg_equal_iff_equal:
assumes x: "x ∈ carrier R" and y: "y ∈ carrier R"
shows "⊖ x = ⊖ y ⟷ x = y"
proof
assume "⊖ x = ⊖ y"
then have "⊖ (⊖ x) = ⊖ (⊖ y)" by simp
with x y show "x = y" by simp
next
assume "x = y"
then show "⊖ x = ⊖ y" by simp
qed
lemma neg_equal_swap:
assumes x: "x ∈ carrier R" and y: "y ∈ carrier R"
shows "(⊖ x = y) = (x = ⊖ y)"
using assms neg_equal_iff_equal [of x "⊖ y"]
by simp
lemma mult2: "x ∈ carrier R ⟹ x ⊕ x = «2» ⊗ x"
by (simp add: of_int_2 l_distr)
end
lemma (in cring) of_int_power [simp]: "«i ^ n» = «i» [^] n"
by (induct n) (simp_all add: m_ac)
definition cring_class_ops :: "'a::comm_ring_1 ring"
where "cring_class_ops ≡ ⦇carrier = UNIV, mult = (*), one = 1, zero = 0, add = (+)⦈"
lemma cring_class: "cring cring_class_ops"
apply unfold_locales
apply (auto simp add: cring_class_ops_def ring_distribs Units_def)
apply (rule_tac x="- x" in exI)
apply simp
done
lemma carrier_class: "x ∈ carrier cring_class_ops"
by (simp add: cring_class_ops_def)
lemma zero_class: "𝟬⇘cring_class_ops⇙ = 0"
by (simp add: cring_class_ops_def)
lemma one_class: "𝟭⇘cring_class_ops⇙ = 1"
by (simp add: cring_class_ops_def)
lemma plus_class: "x ⊕⇘cring_class_ops⇙ y = x + y"
by (simp add: cring_class_ops_def)
lemma times_class: "x ⊗⇘cring_class_ops⇙ y = x * y"
by (simp add: cring_class_ops_def)
lemma uminus_class: "⊖⇘cring_class_ops⇙ x = - x"
apply (simp add: a_inv_def m_inv_def cring_class_ops_def)
apply (rule the_equality)
apply (simp_all add: eq_neg_iff_add_eq_0)
done
lemma minus_class: "x ⊖⇘cring_class_ops⇙ y = x - y"
by (simp add: a_minus_def carrier_class plus_class uminus_class)
lemma power_class: "x [^]⇘cring_class_ops⇙ n = x ^ n"
by (induct n) (simp_all add: one_class times_class
monoid.nat_pow_0 [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]]
monoid.nat_pow_Suc [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]])
lemma of_nat_class: "«n»⇩ℕ⇘cring_class_ops⇙ = of_nat n"
by (induct n) (simp_all add: cring_class_ops_def of_natural_def)
lemma of_int_class: "«i»⇘cring_class_ops⇙ = of_int i"
by (simp add: of_integer_def of_nat_class uminus_class)
lemmas class_simps = zero_class one_class plus_class minus_class uminus_class
times_class power_class of_nat_class of_int_class carrier_class
interpretation cring_class: cring "cring_class_ops::'a::comm_ring_1 ring"
rewrites "(𝟬⇘cring_class_ops⇙::'a) = 0"
and "(𝟭⇘cring_class_ops⇙::'a) = 1"
and "(x::'a) ⊕⇘cring_class_ops⇙ y = x + y"
and "(x::'a) ⊗⇘cring_class_ops⇙ y = x * y"
and "⊖⇘cring_class_ops⇙ (x::'a) = - x"
and "(x::'a) ⊖⇘cring_class_ops⇙ y = x - y"
and "(x::'a) [^]⇘cring_class_ops⇙ n = x ^ n"
and "«n»⇩ℕ⇘cring_class_ops⇙ = of_nat n"
and "((«i»⇘cring_class_ops⇙)::'a) = of_int i"
and "(Int.of_int (numeral m)::'a) = numeral m"