switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
--- a/src/HOL/Library/Commutative_Ring.thy Fri Apr 20 11:21:39 2007 +0200
+++ b/src/HOL/Library/Commutative_Ring.thy Fri Apr 20 11:21:40 2007 +0200
@@ -28,22 +28,22 @@
text {* Interpretation functions for the shadow syntax. *}
-consts
+fun
Ipol :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a pol \<Rightarrow> 'a"
- Ipolex :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a polex \<Rightarrow> 'a"
+where
+ "Ipol l (Pc c) = c"
+ | "Ipol l (Pinj i P) = Ipol (drop i l) P"
+ | "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q"
-primrec
- "Ipol l (Pc c) = c"
- "Ipol l (Pinj i P) = Ipol (drop i l) P"
- "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q"
-
-primrec
- "Ipolex l (Pol P) = Ipol l P"
- "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q"
- "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q"
- "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q"
- "Ipolex l (Pow p n) = Ipolex l p ^ n"
- "Ipolex l (Neg P) = - Ipolex l P"
+fun
+ Ipolex :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a polex \<Rightarrow> 'a"
+where
+ "Ipolex l (Pol P) = Ipol l P"
+ | "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q"
+ | "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q"
+ | "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q"
+ | "Ipolex l (Pow p n) = Ipolex l p ^ n"
+ | "Ipolex l (Neg P) = - Ipolex l P"
text {* Create polynomial normalized polynomials given normalized inputs. *}
@@ -63,122 +63,114 @@
text {* Defining the basic ring operations on normalized polynomials *}
-consts
- add :: "'a::{comm_ring,recpower} pol \<times> 'a pol \<Rightarrow> 'a pol"
- mul :: "'a::{comm_ring,recpower} pol \<times> 'a pol \<Rightarrow> 'a pol"
- neg :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
- sqr :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
- pow :: "'a::{comm_ring,recpower} pol \<times> nat \<Rightarrow> 'a pol"
+function
+ add :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<oplus>" 65)
+where
+ "Pc a \<oplus> Pc b = Pc (a + b)"
+ | "Pc c \<oplus> Pinj i P = Pinj i (P \<oplus> Pc c)"
+ | "Pinj i P \<oplus> Pc c = Pinj i (P \<oplus> Pc c)"
+ | "Pc c \<oplus> PX P i Q = PX P i (Q \<oplus> Pc c)"
+ | "PX P i Q \<oplus> Pc c = PX P i (Q \<oplus> Pc c)"
+ | "Pinj x P \<oplus> Pinj y Q =
+ (if x = y then mkPinj x (P \<oplus> Q)
+ else (if x > y then mkPinj y (Pinj (x - y) P \<oplus> Q)
+ else mkPinj x (Pinj (y - x) Q \<oplus> P)))"
+ | "Pinj x P \<oplus> PX Q y R =
+ (if x = 0 then P \<oplus> PX Q y R
+ else (if x = 1 then PX Q y (R \<oplus> P)
+ else PX Q y (R \<oplus> Pinj (x - 1) P)))"
+ | "PX P x R \<oplus> Pinj y Q =
+ (if y = 0 then PX P x R \<oplus> Q
+ else (if y = 1 then PX P x (R \<oplus> Q)
+ else PX P x (R \<oplus> Pinj (y - 1) Q)))"
+ | "PX P1 x P2 \<oplus> PX Q1 y Q2 =
+ (if x = y then mkPX (P1 \<oplus> Q1) x (P2 \<oplus> Q2)
+ else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<oplus> Q1) y (P2 \<oplus> Q2)
+ else mkPX (PX Q1 (y-x) (Pc 0) \<oplus> P1) x (P2 \<oplus> Q2)))"
+by pat_completeness auto
+termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto
-text {* Addition *}
-recdef add "measure (\<lambda>(x, y). size x + size y)"
- "add (Pc a, Pc b) = Pc (a + b)"
- "add (Pc c, Pinj i P) = Pinj i (add (P, Pc c))"
- "add (Pinj i P, Pc c) = Pinj i (add (P, Pc c))"
- "add (Pc c, PX P i Q) = PX P i (add (Q, Pc c))"
- "add (PX P i Q, Pc c) = PX P i (add (Q, Pc c))"
- "add (Pinj x P, Pinj y Q) =
- (if x=y then mkPinj x (add (P, Q))
- else (if x>y then mkPinj y (add (Pinj (x-y) P, Q))
- else mkPinj x (add (Pinj (y-x) Q, P)) ))"
- "add (Pinj x P, PX Q y R) =
- (if x=0 then add(P, PX Q y R)
- else (if x=1 then PX Q y (add (R, P))
- else PX Q y (add (R, Pinj (x - 1) P))))"
- "add (PX P x R, Pinj y Q) =
- (if y=0 then add(PX P x R, Q)
- else (if y=1 then PX P x (add (R, Q))
- else PX P x (add (R, Pinj (y - 1) Q))))"
- "add (PX P1 x P2, PX Q1 y Q2) =
- (if x=y then mkPX (add (P1, Q1)) x (add (P2, Q2))
- else (if x>y then mkPX (add (PX P1 (x-y) (Pc 0), Q1)) y (add (P2,Q2))
- else mkPX (add (PX Q1 (y-x) (Pc 0), P1)) x (add (P2,Q2)) ))"
-
-text {* Multiplication *}
-recdef mul "measure (\<lambda>(x, y). size x + size y)"
- "mul (Pc a, Pc b) = Pc (a*b)"
- "mul (Pc c, Pinj i P) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))"
- "mul (Pinj i P, Pc c) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))"
- "mul (Pc c, PX P i Q) =
- (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))"
- "mul (PX P i Q, Pc c) =
- (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))"
- "mul (Pinj x P, Pinj y Q) =
- (if x=y then mkPinj x (mul (P, Q))
- else (if x>y then mkPinj y (mul (Pinj (x-y) P, Q))
- else mkPinj x (mul (Pinj (y-x) Q, P)) ))"
- "mul (Pinj x P, PX Q y R) =
- (if x=0 then mul(P, PX Q y R)
- else (if x=1 then mkPX (mul (Pinj x P, Q)) y (mul (R, P))
- else mkPX (mul (Pinj x P, Q)) y (mul (R, Pinj (x - 1) P))))"
- "mul (PX P x R, Pinj y Q) =
- (if y=0 then mul(PX P x R, Q)
- else (if y=1 then mkPX (mul (Pinj y Q, P)) x (mul (R, Q))
- else mkPX (mul (Pinj y Q, P)) x (mul (R, Pinj (y - 1) Q))))"
- "mul (PX P1 x P2, PX Q1 y Q2) =
- add (mkPX (mul (P1, Q1)) (x+y) (mul (P2, Q2)),
- add (mkPX (mul (P1, mkPinj 1 Q2)) x (Pc 0), mkPX (mul (Q1, mkPinj 1 P2)) y (Pc 0)) )"
-(hints simp add: mkPinj_def split: pol.split)
+function
+ mul :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<otimes>" 70)
+where
+ "Pc a \<otimes> Pc b = Pc (a * b)"
+ | "Pc c \<otimes> Pinj i P =
+ (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
+ | "Pinj i P \<otimes> Pc c =
+ (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
+ | "Pc c \<otimes> PX P i Q =
+ (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
+ | "PX P i Q \<otimes> Pc c =
+ (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
+ | "Pinj x P \<otimes> Pinj y Q =
+ (if x = y then mkPinj x (P \<otimes> Q) else
+ (if x > y then mkPinj y (Pinj (x-y) P \<otimes> Q)
+ else mkPinj x (Pinj (y - x) Q \<otimes> P)))"
+ | "Pinj x P \<otimes> PX Q y R =
+ (if x = 0 then P \<otimes> PX Q y R else
+ (if x = 1 then mkPX (Pinj x P \<otimes> Q) y (R \<otimes> P)
+ else mkPX (Pinj x P \<otimes> Q) y (R \<otimes> Pinj (x - 1) P)))"
+ | "PX P x R \<otimes> Pinj y Q =
+ (if y = 0 then PX P x R \<otimes> Q else
+ (if y = 1 then mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Q)
+ else mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Pinj (y - 1) Q)))"
+ | "PX P1 x P2 \<otimes> PX Q1 y Q2 =
+ mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2) \<oplus>
+ (mkPX (P1 \<otimes> mkPinj 1 Q2) x (Pc 0) \<oplus>
+ (mkPX (Q1 \<otimes> mkPinj 1 P2) y (Pc 0)))"
+by pat_completeness auto
+termination by (relation "measure (\<lambda>(x, y). size x + size y)")
+ (auto simp add: mkPinj_def split: pol.split)
text {* Negation*}
-primrec
- "neg (Pc c) = Pc (-c)"
- "neg (Pinj i P) = Pinj i (neg P)"
- "neg (PX P x Q) = PX (neg P) x (neg Q)"
+fun
+ neg :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
+where
+ "neg (Pc c) = Pc (-c)"
+ | "neg (Pinj i P) = Pinj i (neg P)"
+ | "neg (PX P x Q) = PX (neg P) x (neg Q)"
text {* Substraction *}
definition
- sub :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
- "sub p q = add (p, neg q)"
+ sub :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<ominus>" 65)
+where
+ "sub P Q = P \<oplus> neg Q"
text {* Square for Fast Exponentation *}
-primrec
- "sqr (Pc c) = Pc (c * c)"
- "sqr (Pinj i P) = mkPinj i (sqr P)"
- "sqr (PX A x B) = add (mkPX (sqr A) (x + x) (sqr B),
- mkPX (mul (mul (Pc (1 + 1), A), mkPinj 1 B)) x (Pc 0))"
+fun
+ sqr :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
+where
+ "sqr (Pc c) = Pc (c * c)"
+ | "sqr (Pinj i P) = mkPinj i (sqr P)"
+ | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \<oplus>
+ mkPX (Pc (1 + 1) \<otimes> A \<otimes> mkPinj 1 B) x (Pc 0)"
text {* Fast Exponentation *}
-lemma pow_wf: "odd n \<Longrightarrow> (n::nat) div 2 < n" by (cases n) auto
-recdef pow "measure (\<lambda>(x, y). y)"
- "pow (p, 0) = Pc 1"
- "pow (p, n) = (if even n then (pow (sqr p, n div 2)) else mul (p, pow (sqr p, n div 2)))"
-(hints simp add: pow_wf)
-
+fun
+ pow :: "nat \<Rightarrow> 'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol"
+where
+ "pow 0 P = Pc 1"
+ | "pow n P = (if even n then pow (n div 2) (sqr P)
+ else P \<otimes> pow (n div 2) (sqr P))"
+
lemma pow_if:
- "pow (p,n) =
- (if n = 0 then Pc 1 else if even n then pow (sqr p, n div 2)
- else mul (p, pow (sqr p, n div 2)))"
+ "pow n P =
+ (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P)
+ else P \<otimes> pow (n div 2) (sqr P))"
by (cases n) simp_all
-(*
-lemma number_of_nat_B0: "(number_of (w BIT bit.B0) ::nat) = 2* (number_of w)"
-by simp
-
-lemma number_of_nat_even: "even (number_of (w BIT bit.B0)::nat)"
-by simp
-
-lemma pow_even : "pow (p, number_of(w BIT bit.B0)) = pow (sqr p, number_of w)"
- ( is "pow(?p,?n) = pow (_,?n2)")
-proof-
- have "even ?n" by simp
- hence "pow (p, ?n) = pow (sqr p, ?n div 2)"
- apply simp
- apply (cases "IntDef.neg (number_of w)")
- apply simp
- done
-*)
text {* Normalization of polynomial expressions *}
-consts norm :: "'a::{comm_ring,recpower} polex \<Rightarrow> 'a pol"
-primrec
- "norm (Pol P) = P"
- "norm (Add P Q) = add (norm P, norm Q)"
- "norm (Sub p q) = sub (norm p) (norm q)"
- "norm (Mul P Q) = mul (norm P, norm Q)"
- "norm (Pow p n) = pow (norm p, n)"
- "norm (Neg P) = neg (norm P)"
+fun
+ norm :: "'a::{comm_ring,recpower} polex \<Rightarrow> 'a pol"
+where
+ "norm (Pol P) = P"
+ | "norm (Add P Q) = norm P \<oplus> norm Q"
+ | "norm (Sub P Q) = norm P \<ominus> norm Q"
+ | "norm (Mul P Q) = norm P \<otimes> norm Q"
+ | "norm (Pow P n) = pow n (norm P)"
+ | "norm (Neg P) = neg (norm P)"
text {* mkPinj preserve semantics *}
lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)"
@@ -195,7 +187,7 @@
by (induct P arbitrary: l) auto
text {* Addition *}
-lemma add_ci: "Ipol l (add (P, Q)) = Ipol l P + Ipol l Q"
+lemma add_ci: "Ipol l (P \<oplus> Q) = Ipol l P + Ipol l Q"
proof (induct P Q arbitrary: l rule: add.induct)
case (6 x P y Q)
show ?case
@@ -245,26 +237,25 @@
qed (auto simp add: ring_eq_simps)
text {* Multiplication *}
-lemma mul_ci: "Ipol l (mul (P, Q)) = Ipol l P * Ipol l Q"
+lemma mul_ci: "Ipol l (P \<otimes> Q) = Ipol l P * Ipol l Q"
by (induct P Q arbitrary: l rule: mul.induct)
(simp_all add: mkPX_ci mkPinj_ci ring_eq_simps add_ci power_add)
text {* Substraction *}
-lemma sub_ci: "Ipol l (sub p q) = Ipol l p - Ipol l q"
+lemma sub_ci: "Ipol l (P \<ominus> Q) = Ipol l P - Ipol l Q"
by (simp add: add_ci neg_ci sub_def)
text {* Square *}
-lemma sqr_ci: "Ipol ls (sqr p) = Ipol ls p * Ipol ls p"
- by (induct p arbitrary: ls)
+lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P"
+ by (induct P arbitrary: ls)
(simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci ring_eq_simps power_add)
-
text {* Power *}
-lemma even_pow:"even n \<Longrightarrow> pow (p, n) = pow (sqr p, n div 2)"
+lemma even_pow:"even n \<Longrightarrow> pow n P = pow (n div 2) (sqr P)"
by (induct n) simp_all
-lemma pow_ci: "Ipol ls (pow (p, n)) = Ipol ls p ^ n"
-proof (induct n arbitrary: p rule: nat_less_induct)
+lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n"
+proof (induct n arbitrary: P rule: nat_less_induct)
case (1 k)
show ?case
proof (cases k)
@@ -279,7 +270,7 @@
by (simp add: nat_number even_nat_plus_one_div_two)
moreover
from Suc have "l < k" by simp
- with 1 have "\<And>p. Ipol ls (pow (p, l)) = Ipol ls p ^ l" by simp
+ with 1 have "\<And>P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp
moreover
note Suc `even l` even_nat_plus_one_div_two
ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow)
@@ -287,7 +278,7 @@
assume "odd l"
{
fix p
- have "Ipol ls (sqr p) ^ (Suc l div 2) = Ipol ls p ^ Suc l"
+ have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l"
proof (cases l)
case 0
with `odd l` show ?thesis by simp
@@ -296,9 +287,9 @@
with `odd l` have "even w" by simp
have two_times: "2 * (w div 2) = w"
by (simp only: numerals even_nat_div_two_times_two [OF `even w`])
- have "Ipol ls p * Ipol ls p = Ipol ls p ^ Suc (Suc 0)"
+ have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)"
by (simp add: power_Suc)
- then have "Ipol ls p * Ipol ls p = Ipol ls p ^ 2"
+ then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2"
by (simp add: numerals)
with Suc show ?thesis
by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci)
--- a/src/HOL/ex/Commutative_Ring_Complete.thy Fri Apr 20 11:21:39 2007 +0200
+++ b/src/HOL/ex/Commutative_Ring_Complete.thy Fri Apr 20 11:21:40 2007 +0200
@@ -11,19 +11,20 @@
theory Commutative_Ring_Complete
imports Commutative_Ring
begin
-
- (* Fromalization of normal form *)
-consts isnorm :: "('a::{comm_ring,recpower}) pol \<Rightarrow> bool"
-recdef isnorm "measure size"
- "isnorm (Pc c) = True"
- "isnorm (Pinj i (Pc c)) = False"
- "isnorm (Pinj i (Pinj j Q)) = False"
- "isnorm (Pinj 0 P) = False"
- "isnorm (Pinj i (PX Q1 j Q2)) = isnorm (PX Q1 j Q2)"
- "isnorm (PX P 0 Q) = False"
- "isnorm (PX (Pc c) i Q) = (c \<noteq> 0 & isnorm Q)"
- "isnorm (PX (PX P1 j (Pc c)) i Q) = (c\<noteq>0 \<and> isnorm(PX P1 j (Pc c))\<and>isnorm Q)"
- "isnorm (PX P i Q) = (isnorm P \<and> isnorm Q)"
+
+text {* Formalization of normal form *}
+fun
+ isnorm :: "('a::{comm_ring,recpower}) pol \<Rightarrow> bool"
+where
+ "isnorm (Pc c) \<longleftrightarrow> True"
+ | "isnorm (Pinj i (Pc c)) \<longleftrightarrow> False"
+ | "isnorm (Pinj i (Pinj j Q)) \<longleftrightarrow> False"
+ | "isnorm (Pinj 0 P) \<longleftrightarrow> False"
+ | "isnorm (Pinj i (PX Q1 j Q2)) \<longleftrightarrow> isnorm (PX Q1 j Q2)"
+ | "isnorm (PX P 0 Q) \<longleftrightarrow> False"
+ | "isnorm (PX (Pc c) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm Q"
+ | "isnorm (PX (PX P1 j (Pc c)) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm (PX P1 j (Pc c)) \<and> isnorm Q"
+ | "isnorm (PX P i Q) \<longleftrightarrow> isnorm P \<and> isnorm Q"
(* Some helpful lemmas *)
lemma norm_Pinj_0_False:"isnorm (Pinj 0 P) = False"
@@ -57,7 +58,6 @@
case Pinj from prems show ?thesis by(cases x, auto)
qed
-
lemma norm_PXtrans2: assumes A:"isnorm (PX P x Q)" and "isnorm Q2" shows "isnorm (PX P (Suc (n+x)) Q2)"
proof(cases P)
case (PX p1 y p2)
@@ -70,7 +70,7 @@
from prems show ?thesis by(cases x, auto)
qed
- (* mkPX conserves normalizedness (_cn)*)
+text {* mkPX conserves normalizedness (_cn) *}
lemma mkPX_cn:
assumes "x \<noteq> 0" and "isnorm P" and "isnorm Q"
shows "isnorm (mkPX P x Q)"
@@ -87,8 +87,8 @@
with prems Y0 show ?thesis by (cases x, auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto)
qed
- (* add conserves normalizedness *)
-lemma add_cn:"\<lbrakk>isnorm P; (isnorm Q)\<rbrakk> \<Longrightarrow> isnorm (add (P, Q))"
+text {* add conserves normalizedness *}
+lemma add_cn:"isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<oplus> Q)"
proof(induct P Q rule: add.induct)
case (2 c i P2) thus ?case by (cases P2, simp_all, cases i, simp_all)
next
@@ -143,32 +143,32 @@
moreover
{ assume "x=1"
from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
- with prems have "isnorm (add (R, P2))" by simp
+ with prems have "isnorm (R \<oplus> P2)" by simp
with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
moreover
{ assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
then obtain d where X:"x=Suc (Suc d)" ..
from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto)
- with prems NR have "isnorm( add (R, Pinj (x - 1) P2))" "isnorm(PX Q2 y R)" by simp
+ with prems NR have "isnorm (R \<oplus> Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp
with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
ultimately show ?case by blast
next
case (8 Q2 y R x P2)
- have "x=0 \<or> (x = 1) \<or> (x > 1)" by arith
+ have "x = 0 \<or> x = 1 \<or> x > 1" by arith
moreover
{ assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)}
moreover
{ assume "x=1"
from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
- with prems have "isnorm (add (R, P2))" by simp
+ with prems have "isnorm (R \<oplus> P2)" by simp
with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
moreover
{ assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
then obtain d where X:"x=Suc (Suc d)" ..
from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto)
- with prems NR have "isnorm( add (R, Pinj (x - 1) P2))" "isnorm(PX Q2 y R)" by simp
+ with prems NR have "isnorm (R \<oplus> Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp
with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
ultimately show ?case by blast
next
@@ -190,11 +190,11 @@
next case Pc from prems show ?thesis by(cases d, auto)
next case Pinj from prems show ?thesis by(cases d, auto)
qed
- ultimately have "isnorm (add (P2, Q2))" "isnorm (add (PX P1 (x - y) (Pc 0), Q1))" by auto
+ ultimately have "isnorm (P2 \<oplus> Q2)" "isnorm (PX P1 (x - y) (Pc 0) \<oplus> Q1)" by auto
with Y0 sm1 sm2 have ?case by (simp add: mkPX_cn)}
moreover
{assume "x=y"
- from prems NP1 NP2 NQ1 NQ2 have "isnorm (add (P2, Q2))" "isnorm (add (P1, Q1))" by auto
+ from prems NP1 NP2 NQ1 NQ2 have "isnorm (P2 \<oplus> Q2)" "isnorm (P1 \<oplus> Q1)" by auto
with Y0 prems have ?case by (simp add: mkPX_cn) }
moreover
{assume sm1:"x<y" hence "EX d. y=d+x" by arith
@@ -208,13 +208,13 @@
next case Pc from prems show ?thesis by(cases d, auto)
next case Pinj from prems show ?thesis by(cases d, auto)
qed
- ultimately have "isnorm (add (P2, Q2))" "isnorm (add (PX Q1 (y - x) (Pc 0), P1))" by auto
+ ultimately have "isnorm (P2 \<oplus> Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \<oplus> P1)" by auto
with X0 sm1 sm2 have ?case by (simp add: mkPX_cn)}
ultimately show ?case by blast
-qed(simp)
+qed simp
- (* mul concerves normalizedness *)
-lemma mul_cn :"\<lbrakk>isnorm P; (isnorm Q)\<rbrakk> \<Longrightarrow> isnorm (mul (P, Q))"
+text {* mul concerves normalizedness *}
+lemma mul_cn :"isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<otimes> Q)"
proof(induct P Q rule: mul.induct)
case (2 c i P2) thus ?case
by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn)
@@ -278,7 +278,7 @@
moreover
{ assume "x=1"
from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
- with prems have "isnorm (mul (R, P2))" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
+ with prems have "isnorm (R \<otimes> P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
with Y0 prems have ?case by (simp add: mkPX_cn)}
moreover
{ assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
@@ -290,7 +290,7 @@
from prems have "isnorm (Pinj x P2)" by(cases P2, auto)
moreover
note prems
- ultimately have "isnorm (mul (R, Pinj (x - 1) P2))" "isnorm (mul (Pinj x P2, Q2))" by auto
+ ultimately have "isnorm (R \<otimes> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<otimes> Q2)" by auto
with Y0 X have ?case by (simp add: mkPX_cn)}
ultimately show ?case by blast
next
@@ -302,7 +302,7 @@
moreover
{ assume "x=1"
from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
- with prems have "isnorm (mul (R, P2))" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
+ with prems have "isnorm (R \<otimes> P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
with Y0 prems have ?case by (simp add: mkPX_cn) }
moreover
{ assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
@@ -314,7 +314,7 @@
from prems have "isnorm (Pinj x P2)" by(cases P2, auto)
moreover
note prems
- ultimately have "isnorm (mul (R, Pinj (x - 1) P2))" "isnorm (mul (Pinj x P2, Q2))" by auto
+ ultimately have "isnorm (R \<otimes> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<otimes> Q2)" by auto
with Y0 X have ?case by (simp add: mkPX_cn) }
ultimately show ?case by blast
next
@@ -326,17 +326,20 @@
from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2])
moreover
from prems have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2])
- ultimately have "isnorm (mul (P1, Q1))" "isnorm (mul (P2, Q2))" "isnorm (mul (P1, mkPinj 1 Q2))" "isnorm (mul (Q1, mkPinj 1 P2))"
+ ultimately have "isnorm (P1 \<otimes> Q1)" "isnorm (P2 \<otimes> Q2)"
+ "isnorm (P1 \<otimes> mkPinj 1 Q2)" "isnorm (Q1 \<otimes> mkPinj 1 P2)"
by (auto simp add: mkPinj_cn)
- with prems X0 Y0 have "isnorm (mkPX (mul (P1, Q1)) (x + y) (mul (P2, Q2)))" "isnorm (mkPX (mul (P1, mkPinj (Suc 0) Q2)) x (Pc 0))"
- "isnorm (mkPX (mul (Q1, mkPinj (Suc 0) P2)) y (Pc 0))"
+ with prems X0 Y0 have
+ "isnorm (mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2))"
+ "isnorm (mkPX (P1 \<otimes> mkPinj (Suc 0) Q2) x (Pc 0))"
+ "isnorm (mkPX (Q1 \<otimes> mkPinj (Suc 0) P2) y (Pc 0))"
by (auto simp add: mkPX_cn)
thus ?case by (simp add: add_cn)
qed(simp)
- (* neg conserves normalizedness *)
+text {* neg conserves normalizedness *}
lemma neg_cn: "isnorm P \<Longrightarrow> isnorm (neg P)"
-proof(induct P rule: neg.induct)
+proof (induct P)
case (Pinj i P2)
from prems have "isnorm P2" by (simp add: norm_Pinj[of i P2])
with prems show ?case by(cases P2, auto, cases i, auto)
@@ -353,11 +356,11 @@
qed(cases x, auto)
qed(simp)
- (* sub conserves normalizedness *)
-lemma sub_cn:"\<lbrakk>isnorm p; isnorm q\<rbrakk> \<Longrightarrow> isnorm (sub p q)"
+text {* sub conserves normalizedness *}
+lemma sub_cn:"isnorm p \<Longrightarrow> isnorm q \<Longrightarrow> isnorm (p \<ominus> q)"
by (simp add: sub_def add_cn neg_cn)
- (* sqr conserves normalizizedness *)
+text {* sqr conserves normalizizedness *}
lemma sqr_cn:"isnorm P \<Longrightarrow> isnorm (sqr P)"
proof(induct P)
case (Pinj i Q)
@@ -365,23 +368,25 @@
next
case (PX P1 x P2)
from prems have "x+x~=0" "isnorm P2" "isnorm P1" by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2])
- with prems have "isnorm (mkPX (mul (mul (Pc ((1\<Colon>'a) + (1\<Colon>'a)), P1), mkPinj (Suc 0) P2)) x (Pc (0\<Colon>'a)))"
- and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" by( auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
- thus ?case by( auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
-qed(simp)
+ with prems have
+ "isnorm (mkPX (Pc (1 + 1) \<otimes> P1 \<otimes> mkPinj (Suc 0) P2) x (Pc 0))"
+ and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))"
+ by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
+ thus ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
+qed simp
-
- (* pow conserves normalizedness *)
-lemma pow_cn:"!! P. \<lbrakk>isnorm P\<rbrakk> \<Longrightarrow> isnorm (pow (P, n))"
-proof(induct n rule: nat_less_induct)
+text {* pow conserves normalizedness *}
+lemma pow_cn:"isnorm P \<Longrightarrow> isnorm (pow n P)"
+proof (induct n arbitrary: P rule: nat_less_induct)
case (1 k)
show ?case
- proof(cases "k=0")
+ proof (cases "k=0")
case False
- hence K2:"k div 2 < k" by (cases k, auto)
+ then have K2:"k div 2 < k" by (cases k, auto)
from prems have "isnorm (sqr P)" by (simp add: sqr_cn)
- with prems K2 show ?thesis by(simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn)
- qed(simp)
+ with prems K2 show ?thesis
+ by (simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn)
+ qed simp
qed
end