# HG changeset patch # User wenzelm # Date 1329824700 -3600 # Node ID 26977429b784d36da5d8745c47c10fd99082362f # Parent 0133d31f9ab4ca122025f98794d7cc8440a17f58# Parent 092f4eca98485add0d670d01abbde4ca76dd1bf3 merged; diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Mon Feb 20 22:35:32 2012 +0100 +++ b/src/HOL/Algebra/Group.thy Tue Feb 21 12:45:00 2012 +0100 @@ -41,7 +41,7 @@ begin definition "int_pow G a z = (let p = nat_rec \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) - in if neg z then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))" + in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))" end locale monoid = @@ -391,7 +391,7 @@ text {* Power *} lemma (in group) int_pow_def2: - "a (^) (z::int) = (if neg z then inv (a (^) (nat (-z))) else a (^) (nat z))" + "a (^) (z::int) = (if z < 0 then inv (a (^) (nat (-z))) else a (^) (nat z))" by (simp add: int_pow_def nat_pow_def Let_def) lemma (in group) int_pow_0 [simp]: diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/Big_Operators.thy diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/Complete_Lattices.thy diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Mon Feb 20 22:35:32 2012 +0100 +++ b/src/HOL/Divides.thy Tue Feb 21 12:45:00 2012 +0100 @@ -523,9 +523,6 @@ means of @{const divmod_nat_rel}: *} -instantiation nat :: semiring_div -begin - definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (THE qr. divmod_nat_rel m n qr)" @@ -543,28 +540,39 @@ shows "divmod_nat m n = qr" using assms by (auto intro: divmod_nat_rel_unique divmod_nat_rel_divmod_nat) +instantiation nat :: semiring_div +begin + definition div_nat where "m div n = fst (divmod_nat m n)" +lemma fst_divmod_nat [simp]: + "fst (divmod_nat m n) = m div n" + by (simp add: div_nat_def) + definition mod_nat where "m mod n = snd (divmod_nat m n)" +lemma snd_divmod_nat [simp]: + "snd (divmod_nat m n) = m mod n" + by (simp add: mod_nat_def) + lemma divmod_nat_div_mod: "divmod_nat m n = (m div n, m mod n)" - unfolding div_nat_def mod_nat_def by simp + by (simp add: prod_eq_iff) lemma div_eq: assumes "divmod_nat_rel m n (q, r)" shows "m div n = q" - using assms by (auto dest: divmod_nat_eq simp add: divmod_nat_div_mod) + using assms by (auto dest!: divmod_nat_eq simp add: prod_eq_iff) lemma mod_eq: assumes "divmod_nat_rel m n (q, r)" shows "m mod n = r" - using assms by (auto dest: divmod_nat_eq simp add: divmod_nat_div_mod) + using assms by (auto dest!: divmod_nat_eq simp add: prod_eq_iff) lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)" - by (simp add: div_nat_def mod_nat_def divmod_nat_rel_divmod_nat) + using divmod_nat_rel_divmod_nat by (simp add: divmod_nat_div_mod) lemma divmod_nat_zero: "divmod_nat m 0 = (0, m)" @@ -613,25 +621,25 @@ fixes m n :: nat assumes "m < n" shows "m div n = 0" - using assms divmod_nat_base divmod_nat_div_mod by simp + using assms divmod_nat_base by (simp add: prod_eq_iff) lemma le_div_geq: fixes m n :: nat assumes "0 < n" and "n \ m" shows "m div n = Suc ((m - n) div n)" - using assms divmod_nat_step divmod_nat_div_mod by simp + using assms divmod_nat_step by (simp add: prod_eq_iff) lemma mod_less [simp]: fixes m n :: nat assumes "m < n" shows "m mod n = m" - using assms divmod_nat_base divmod_nat_div_mod by simp + using assms divmod_nat_base by (simp add: prod_eq_iff) lemma le_mod_geq: fixes m n :: nat assumes "n \ m" shows "m mod n = (m - n) mod n" - using assms divmod_nat_step divmod_nat_div_mod by (cases "n = 0") simp_all + using assms divmod_nat_step by (cases "n = 0") (simp_all add: prod_eq_iff) instance proof - have [simp]: "\n::nat. n div 0 = 0" @@ -673,8 +681,7 @@ lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = divmod_nat (m - n) n in (Suc q, r))" -by (simp add: divmod_nat_zero divmod_nat_base divmod_nat_step) - (simp add: divmod_nat_div_mod) + by (simp add: prod_eq_iff prod_case_beta not_less le_div_geq le_mod_geq) text {* Simproc for cancelling @{const div} and @{const mod} *} @@ -760,27 +767,23 @@ subsubsection {* Quotient and Remainder *} lemma divmod_nat_rel_mult1_eq: - "divmod_nat_rel b c (q, r) \ c > 0 + "divmod_nat_rel b c (q, r) \ divmod_nat_rel (a * b) c (a * q + a * r div c, a * r mod c)" by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) lemma div_mult1_eq: "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)" -apply (cases "c = 0", simp) -apply (blast intro: divmod_nat_rel [THEN divmod_nat_rel_mult1_eq, THEN div_eq]) -done +by (blast intro: divmod_nat_rel [THEN divmod_nat_rel_mult1_eq, THEN div_eq]) lemma divmod_nat_rel_add1_eq: - "divmod_nat_rel a c (aq, ar) \ divmod_nat_rel b c (bq, br) \ c > 0 + "divmod_nat_rel a c (aq, ar) \ divmod_nat_rel b c (bq, br) \ divmod_nat_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)" by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) (*NOT suitable for rewriting: the RHS has an instance of the LHS*) lemma div_add1_eq: "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" -apply (cases "c = 0", simp) -apply (blast intro: divmod_nat_rel_add1_eq [THEN div_eq] divmod_nat_rel) -done +by (blast intro: divmod_nat_rel_add1_eq [THEN div_eq] divmod_nat_rel) lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" apply (cut_tac m = q and n = c in mod_less_divisor) @@ -790,29 +793,22 @@ done lemma divmod_nat_rel_mult2_eq: - "divmod_nat_rel a b (q, r) \ 0 < b \ 0 < c + "divmod_nat_rel a b (q, r) \ divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)" by (auto simp add: mult_ac divmod_nat_rel_def add_mult_distrib2 [symmetric] mod_lemma) lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" - apply (cases "b = 0", simp) - apply (cases "c = 0", simp) - apply (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_eq]) - done +by (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_eq]) lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)" - apply (cases "b = 0", simp) - apply (cases "c = 0", simp) - apply (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_eq]) - done - - -subsubsection{*Further Facts about Quotient and Remainder*} +by (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_eq]) + + +subsubsection {* Further Facts about Quotient and Remainder *} lemma div_1 [simp]: "m div Suc 0 = m" by (induct m) (simp_all add: div_geq) - (* Monotonicity of div in first argument *) lemma div_le_mono [rule_format (no_asm)]: "\m::nat. m \ n --> (m div k) \ (n div k)" @@ -1018,7 +1014,7 @@ qed -subsubsection {*An ``induction'' law for modulus arithmetic.*} +subsubsection {* An ``induction'' law for modulus arithmetic. *} lemma mod_induct_0: assumes step: "\i P ((Suc i) mod p)" @@ -1211,8 +1207,6 @@ (auto simp add: mult_2) text{*algorithm for the general case @{term "b\0"}*} -definition negateSnd :: "int \ int \ int \ int" where - [code_unfold]: "negateSnd = apsnd uminus" definition divmod_int :: "int \ int \ int \ int" where --{*The full division algorithm considers all possible signs for a, b @@ -1220,19 +1214,27 @@ @{term negDivAlg} requires @{term "a<0"}.*} "divmod_int a b = (if 0 \ a then if 0 \ b then posDivAlg a b else if a = 0 then (0, 0) - else negateSnd (negDivAlg (-a) (-b)) + else apsnd uminus (negDivAlg (-a) (-b)) else if 0 < b then negDivAlg a b - else negateSnd (posDivAlg (-a) (-b)))" + else apsnd uminus (posDivAlg (-a) (-b)))" instantiation int :: Divides.div begin -definition +definition div_int where "a div b = fst (divmod_int a b)" -definition - "a mod b = snd (divmod_int a b)" +lemma fst_divmod_int [simp]: + "fst (divmod_int a b) = a div b" + by (simp add: div_int_def) + +definition mod_int where + "a mod b = snd (divmod_int a b)" + +lemma snd_divmod_int [simp]: + "snd (divmod_int a b) = a mod b" + by (simp add: mod_int_def) instance .. @@ -1240,7 +1242,7 @@ lemma divmod_int_mod_div: "divmod_int p q = (p div q, p mod q)" - by (auto simp add: div_int_def mod_int_def) + by (simp add: prod_eq_iff) text{* Here is the division algorithm in ML: @@ -1271,8 +1273,7 @@ *} - -subsubsection{*Uniqueness and Monotonicity of Quotients and Remainders*} +subsubsection {* Uniqueness and Monotonicity of Quotients and Remainders *} lemma unique_quotient_lemma: "[| b*q' + r' \ b*q + r; 0 \ r'; r' < b; r < b |] @@ -1294,7 +1295,7 @@ auto) lemma unique_quotient: - "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] + "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r') |] ==> q = q'" apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm) apply (blast intro: order_antisym @@ -1304,7 +1305,7 @@ lemma unique_remainder: - "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] + "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r') |] ==> r = r'" apply (subgoal_tac "q = q'") apply (simp add: divmod_int_rel_def) @@ -1312,7 +1313,7 @@ done -subsubsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*} +subsubsection {* Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends *} text{*And positive divisors*} @@ -1347,7 +1348,7 @@ done -subsubsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*} +subsubsection {* Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends *} text{*And positive divisors*} @@ -1377,7 +1378,7 @@ done -subsubsection{*Existence Shown by Proving the Division Algorithm to be Correct*} +subsubsection {* Existence Shown by Proving the Division Algorithm to be Correct *} (*the case a=0*) lemma divmod_int_rel_0: "b \ 0 ==> divmod_int_rel 0 b (0, 0)" @@ -1389,10 +1390,7 @@ lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" by (subst negDivAlg.simps, auto) -lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" -by (simp add: negateSnd_def) - -lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (negateSnd qr)" +lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (apsnd uminus qr)" by (auto simp add: split_ifs divmod_int_rel_def) lemma divmod_int_correct: "b \ 0 ==> divmod_int_rel a b (divmod_int a b)" @@ -1411,7 +1409,7 @@ lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" apply (case_tac "b = 0", simp) apply (cut_tac a = a and b = b in divmod_int_correct) -apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) +apply (auto simp add: divmod_int_rel_def prod_eq_iff) done lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" @@ -1442,7 +1440,7 @@ lemma pos_mod_conj : "(0::int) < b ==> 0 \ a mod b & a mod b < b" apply (cut_tac a = a and b = b in divmod_int_correct) -apply (auto simp add: divmod_int_rel_def mod_int_def) +apply (auto simp add: divmod_int_rel_def prod_eq_iff) done lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1] @@ -1450,24 +1448,28 @@ lemma neg_mod_conj : "b < (0::int) ==> a mod b \ 0 & b < a mod b" apply (cut_tac a = a and b = b in divmod_int_correct) -apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) +apply (auto simp add: divmod_int_rel_def prod_eq_iff) done lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1] and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2] -subsubsection{*General Properties of div and mod*} +subsubsection {* General Properties of div and mod *} lemma divmod_int_rel_div_mod: "b \ 0 ==> divmod_int_rel a b (a div b, a mod b)" apply (cut_tac a = a and b = b in zmod_zdiv_equality) apply (force simp add: divmod_int_rel_def linorder_neq_iff) done -lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a div b = q" +lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r) |] ==> a div b = q" +apply (cases "b = 0") +apply (simp add: divmod_int_rel_def) by (simp add: divmod_int_rel_div_mod [THEN unique_quotient]) -lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a mod b = r" +lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r) |] ==> a mod b = r" +apply (cases "b = 0") +apply (simp add: divmod_int_rel_def) by (simp add: divmod_int_rel_div_mod [THEN unique_remainder]) lemma div_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a div b = 0" @@ -1521,7 +1523,7 @@ done -subsubsection{*Laws for div and mod with Unary Minus*} +subsubsection {* Laws for div and mod with Unary Minus *} lemma zminus1_lemma: "divmod_int_rel a b (q, r) @@ -1569,7 +1571,7 @@ unfolding zmod_zminus2_eq_if by auto -subsubsection{*Division of a Number by Itself*} +subsubsection {* Division of a Number by Itself *} lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \ q" apply (subgoal_tac "0 < a*q") @@ -1582,7 +1584,7 @@ apply (simp add: right_diff_distrib) done -lemma self_quotient: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> q = 1" +lemma self_quotient: "[| divmod_int_rel a a (q, r) |] ==> q = 1" apply (simp add: split_ifs divmod_int_rel_def linorder_neq_iff) apply (rule order_antisym, safe, simp_all) apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) @@ -1590,8 +1592,8 @@ apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ done -lemma self_remainder: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> r = 0" -apply (frule self_quotient, assumption) +lemma self_remainder: "[| divmod_int_rel a a (q, r) |] ==> r = 0" +apply (frule self_quotient) apply (simp add: divmod_int_rel_def) done @@ -1605,7 +1607,7 @@ done -subsubsection{*Computation of Division and Remainder*} +subsubsection {* Computation of Division and Remainder *} lemma zdiv_zero [simp]: "(0::int) div b = 0" by (simp add: div_int_def divmod_int_def) @@ -1638,40 +1640,39 @@ text{*a positive, b negative *} lemma div_pos_neg: - "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" + "[| 0 < a; b < 0 |] ==> a div b = fst (apsnd uminus (negDivAlg (-a) (-b)))" by (simp add: div_int_def divmod_int_def) lemma mod_pos_neg: - "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" + "[| 0 < a; b < 0 |] ==> a mod b = snd (apsnd uminus (negDivAlg (-a) (-b)))" by (simp add: mod_int_def divmod_int_def) text{*a negative, b negative *} lemma div_neg_neg: - "[| a < 0; b \ 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" + "[| a < 0; b \ 0 |] ==> a div b = fst (apsnd uminus (posDivAlg (-a) (-b)))" by (simp add: div_int_def divmod_int_def) lemma mod_neg_neg: - "[| a < 0; b \ 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" + "[| a < 0; b \ 0 |] ==> a mod b = snd (apsnd uminus (posDivAlg (-a) (-b)))" by (simp add: mod_int_def divmod_int_def) text {*Simplify expresions in which div and mod combine numerical constants*} lemma int_div_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a div b = q" - by (rule divmod_int_rel_div [of a b q r], - simp add: divmod_int_rel_def, simp) + by (rule divmod_int_rel_div [of a b q r]) (simp add: divmod_int_rel_def) lemma int_div_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a div b = q" by (rule divmod_int_rel_div [of a b q r], - simp add: divmod_int_rel_def, simp) + simp add: divmod_int_rel_def) lemma int_mod_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a mod b = r" by (rule divmod_int_rel_mod [of a b q r], - simp add: divmod_int_rel_def, simp) + simp add: divmod_int_rel_def) lemma int_mod_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a mod b = r" by (rule divmod_int_rel_mod [of a b q r], - simp add: divmod_int_rel_def, simp) + simp add: divmod_int_rel_def) lemmas arithmetic_simps = arith_simps @@ -1748,7 +1749,7 @@ lemmas negDivAlg_eqn_1_number_of [simp] = negDivAlg_eqn [of concl: 1 "number_of w"] for w -subsubsection{*Monotonicity in the First Argument (Dividend)*} +subsubsection {* Monotonicity in the First Argument (Dividend) *} lemma zdiv_mono1: "[| a \ a'; 0 < (b::int) |] ==> a div b \ a' div b" apply (cut_tac a = a and b = b in zmod_zdiv_equality) @@ -1767,7 +1768,7 @@ done -subsubsection{*Monotonicity in the Second Argument (Divisor)*} +subsubsection {* Monotonicity in the Second Argument (Divisor) *} lemma q_pos_lemma: "[| 0 \ b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \ (q'::int)" @@ -1829,12 +1830,12 @@ done -subsubsection{*More Algebraic Laws for div and mod*} +subsubsection {* More Algebraic Laws for div and mod *} text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *} lemma zmult1_lemma: - "[| divmod_int_rel b c (q, r); c \ 0 |] + "[| divmod_int_rel b c (q, r) |] ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac) @@ -1856,7 +1857,7 @@ text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *} lemma zadd1_lemma: - "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br); c \ 0 |] + "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br) |] ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib) @@ -1890,8 +1891,7 @@ done moreover with `c \ 0` divmod_int_rel_div_mod have "divmod_int_rel b c (b div c, b mod c)" by auto ultimately have "divmod_int_rel (a * b) (a * c) (b div c, a * (b mod c))" . - moreover from `a \ 0` `c \ 0` have "a * c \ 0" by simp - ultimately show ?thesis by (rule divmod_int_rel_div) + from this show ?thesis by (rule divmod_int_rel_div) qed qed auto @@ -1905,7 +1905,7 @@ case False with assms posDivAlg_correct have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" by simp - from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] + from divmod_int_rel_div [OF this] divmod_int_rel_mod [OF this] show ?thesis by simp qed @@ -1918,7 +1918,7 @@ from assms negDivAlg_correct have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" by simp - from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] + from divmod_int_rel_div [OF this] divmod_int_rel_mod [OF this] show ?thesis by simp qed @@ -1929,7 +1929,7 @@ lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] -subsubsection{*Proving @{term "a div (b*c) = (a div b) div c"} *} +subsubsection {* Proving @{term "a div (b*c) = (a div b) div c"} *} (*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems @@ -1972,7 +1972,7 @@ apply simp done -lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); b \ 0; 0 < c |] +lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); 0 < c |] ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff zero_less_mult_iff right_distrib [symmetric] @@ -1990,7 +1990,7 @@ done -subsubsection {*Splitting Rules for div and mod*} +subsubsection {* Splitting Rules for div and mod *} text{*The proofs of the two lemmas below are essentially identical*} @@ -2051,7 +2051,7 @@ declare split_zmod [of _ _ "number_of k", arith_split] for k -subsubsection{*Speeding up the Division Algorithm with Shifting*} +subsubsection {* Speeding up the Division Algorithm with Shifting *} text{*computing div by shifting *} @@ -2105,7 +2105,7 @@ done -subsubsection{*Computing mod by Shifting (proofs resemble those for div)*} +subsubsection {* Computing mod by Shifting (proofs resemble those for div) *} lemma pos_zmod_mult_2: fixes a b :: int @@ -2169,7 +2169,7 @@ qed -subsubsection{*Quotients of Signs*} +subsubsection {* Quotients of Signs *} lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" apply (subgoal_tac "a div b \ -1", force) @@ -2225,7 +2225,6 @@ using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b]apply simp done - lemma zmod_le_nonneg_dividend: "(m::int) \ 0 ==> m mod k \ m" apply (rule split_zmod[THEN iffD2]) apply(fastforce dest: q_pos_lemma intro: split_mult_pos_le) @@ -2384,7 +2383,7 @@ proof- from xy have th: "int x - int y = int (x - y)" by simp from xyn have "int x mod int n = int y mod int n" - by (simp add: zmod_int[symmetric]) + by (simp add: zmod_int [symmetric]) hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) hence "n dvd x - y" by (simp add: th zdvd_int) then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Feb 20 22:35:32 2012 +0100 +++ b/src/HOL/IsaMakefile Tue Feb 21 12:45:00 2012 +0100 @@ -1059,9 +1059,9 @@ ex/Lagrange.thy ex/List_to_Set_Comprehension_Examples.thy \ ex/LocaleTest2.thy ex/MT.thy ex/MergeSort.thy ex/Meson_Test.thy \ ex/MonoidGroup.thy ex/Multiquote.thy ex/NatSum.thy \ - ex/Normalization_by_Evaluation.thy ex/Numeral.thy ex/PER.thy \ - ex/PresburgerEx.thy ex/Primrec.thy ex/Quickcheck_Examples.thy \ - ex/Quickcheck_Lattice_Examples.thy \ + ex/Normalization_by_Evaluation.thy ex/Numeral_Representation.thy \ + ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy \ + ex/Quickcheck_Examples.thy ex/Quickcheck_Lattice_Examples.thy \ ex/Quickcheck_Narrowing_Examples.thy ex/Quicksort.thy ex/ROOT.ML \ ex/Records.thy ex/ReflectionEx.thy ex/Refute_Examples.thy \ ex/SAT_Examples.thy ex/Serbian.thy ex/Set_Theory.thy \ diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/Lattices.thy diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/Matrix/ComputeNumeral.thy --- a/src/HOL/Matrix/ComputeNumeral.thy Mon Feb 20 22:35:32 2012 +0100 +++ b/src/HOL/Matrix/ComputeNumeral.thy Tue Feb 21 12:45:00 2012 +0100 @@ -147,19 +147,16 @@ lemma adjust: "adjust b (q, r) = (if 0 \ r - b then (2 * q + 1, r - b) else (2 * q, r))" by (auto simp only: adjust_def) -lemma negateSnd: "negateSnd (q, r) = (q, -r)" - by (simp add: negateSnd_def) - lemma divmod: "divmod_int a b = (if 0\a then if 0\b then posDivAlg a b else if a=0 then (0, 0) - else negateSnd (negDivAlg (-a) (-b)) + else apsnd uminus (negDivAlg (-a) (-b)) else if 0 num" where - "inc One = Dig0 One" -| "inc (Dig0 x) = Dig1 x" -| "inc (Dig1 x) = Dig0 (inc x)" - -text {* Converting between type @{typ num} and type @{typ nat} *} - -primrec nat_of_num :: "num \ nat" where - "nat_of_num One = Suc 0" -| "nat_of_num (Dig0 x) = nat_of_num x + nat_of_num x" -| "nat_of_num (Dig1 x) = Suc (nat_of_num x + nat_of_num x)" - -primrec num_of_nat :: "nat \ num" where - "num_of_nat 0 = One" -| "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" - -lemma nat_of_num_pos: "0 < nat_of_num x" - by (induct x) simp_all - -lemma nat_of_num_neq_0: "nat_of_num x \ 0" - by (induct x) simp_all - -lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" - by (induct x) simp_all - -lemma num_of_nat_double: - "0 < n \ num_of_nat (n + n) = Dig0 (num_of_nat n)" - by (induct n) simp_all - -text {* - Type @{typ num} is isomorphic to the strictly positive - natural numbers. -*} - -lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" - by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) - -lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" - by (induct n) (simp_all add: nat_of_num_inc) - -lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" -proof - assume "nat_of_num x = nat_of_num y" - then have "num_of_nat (nat_of_num x) = num_of_nat (nat_of_num y)" by simp - then show "x = y" by (simp add: nat_of_num_inverse) -qed simp - -lemma num_induct [case_names One inc]: - fixes P :: "num \ bool" - assumes One: "P One" - and inc: "\x. P x \ P (inc x)" - shows "P x" -proof - - obtain n where n: "Suc n = nat_of_num x" - by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0) - have "P (num_of_nat (Suc n))" - proof (induct n) - case 0 show ?case using One by simp - next - case (Suc n) - then have "P (inc (num_of_nat (Suc n)))" by (rule inc) - then show "P (num_of_nat (Suc (Suc n)))" by simp - qed - with n show "P x" - by (simp add: nat_of_num_inverse) -qed - -text {* - From now on, there are two possible models for @{typ num}: as - positive naturals (rule @{text "num_induct"}) and as digit - representation (rules @{text "num.induct"}, @{text "num.cases"}). - - It is not entirely clear in which context it is better to use the - one or the other, or whether the construction should be reversed. -*} - - -subsection {* Numeral operations *} - -ML {* -structure Dig_Simps = Named_Thms -( - val name = @{binding numeral} - val description = "simplification rules for numerals" -) -*} - -setup Dig_Simps.setup - -instantiation num :: "{plus,times,ord}" -begin - -definition plus_num :: "num \ num \ num" where - "m + n = num_of_nat (nat_of_num m + nat_of_num n)" - -definition times_num :: "num \ num \ num" where - "m * n = num_of_nat (nat_of_num m * nat_of_num n)" - -definition less_eq_num :: "num \ num \ bool" where - "m \ n \ nat_of_num m \ nat_of_num n" - -definition less_num :: "num \ num \ bool" where - "m < n \ nat_of_num m < nat_of_num n" - -instance .. - -end - -lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" - unfolding plus_num_def - by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) - -lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" - unfolding times_num_def - by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) - -lemma Dig_plus [numeral, simp, code]: - "One + One = Dig0 One" - "One + Dig0 m = Dig1 m" - "One + Dig1 m = Dig0 (m + One)" - "Dig0 n + One = Dig1 n" - "Dig0 n + Dig0 m = Dig0 (n + m)" - "Dig0 n + Dig1 m = Dig1 (n + m)" - "Dig1 n + One = Dig0 (n + One)" - "Dig1 n + Dig0 m = Dig1 (n + m)" - "Dig1 n + Dig1 m = Dig0 (n + m + One)" - by (simp_all add: num_eq_iff nat_of_num_add) - -lemma Dig_times [numeral, simp, code]: - "One * One = One" - "One * Dig0 n = Dig0 n" - "One * Dig1 n = Dig1 n" - "Dig0 n * One = Dig0 n" - "Dig0 n * Dig0 m = Dig0 (n * Dig0 m)" - "Dig0 n * Dig1 m = Dig0 (n * Dig1 m)" - "Dig1 n * One = Dig1 n" - "Dig1 n * Dig0 m = Dig0 (n * Dig0 m + m)" - "Dig1 n * Dig1 m = Dig1 (n * Dig1 m + m)" - by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult - left_distrib right_distrib) - -lemma less_eq_num_code [numeral, simp, code]: - "One \ n \ True" - "Dig0 m \ One \ False" - "Dig1 m \ One \ False" - "Dig0 m \ Dig0 n \ m \ n" - "Dig0 m \ Dig1 n \ m \ n" - "Dig1 m \ Dig1 n \ m \ n" - "Dig1 m \ Dig0 n \ m < n" - using nat_of_num_pos [of n] nat_of_num_pos [of m] - by (auto simp add: less_eq_num_def less_num_def) - -lemma less_num_code [numeral, simp, code]: - "m < One \ False" - "One < One \ False" - "One < Dig0 n \ True" - "One < Dig1 n \ True" - "Dig0 m < Dig0 n \ m < n" - "Dig0 m < Dig1 n \ m \ n" - "Dig1 m < Dig1 n \ m < n" - "Dig1 m < Dig0 n \ m < n" - using nat_of_num_pos [of n] nat_of_num_pos [of m] - by (auto simp add: less_eq_num_def less_num_def) - -text {* Rules using @{text One} and @{text inc} as constructors *} - -lemma add_One: "x + One = inc x" - by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) - -lemma add_inc: "x + inc y = inc (x + y)" - by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) - -lemma mult_One: "x * One = x" - by (simp add: num_eq_iff nat_of_num_mult) - -lemma mult_inc: "x * inc y = x * y + x" - by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) - -text {* A double-and-decrement function *} - -primrec DigM :: "num \ num" where - "DigM One = One" -| "DigM (Dig0 n) = Dig1 (DigM n)" -| "DigM (Dig1 n) = Dig1 (Dig0 n)" - -lemma DigM_plus_one: "DigM n + One = Dig0 n" - by (induct n) simp_all - -lemma add_One_commute: "One + n = n + One" - by (induct n) simp_all - -lemma one_plus_DigM: "One + DigM n = Dig0 n" - by (simp add: add_One_commute DigM_plus_one) - -text {* Squaring and exponentiation *} - -primrec square :: "num \ num" where - "square One = One" -| "square (Dig0 n) = Dig0 (Dig0 (square n))" -| "square (Dig1 n) = Dig1 (Dig0 (square n + n))" - -primrec pow :: "num \ num \ num" where - "pow x One = x" -| "pow x (Dig0 y) = square (pow x y)" -| "pow x (Dig1 y) = x * square (pow x y)" - - -subsection {* Binary numerals *} - -text {* - We embed binary representations into a generic algebraic - structure using @{text of_num}. -*} - -class semiring_numeral = semiring + monoid_mult -begin - -primrec of_num :: "num \ 'a" where - of_num_One [numeral]: "of_num One = 1" -| "of_num (Dig0 n) = of_num n + of_num n" -| "of_num (Dig1 n) = of_num n + of_num n + 1" - -lemma of_num_inc: "of_num (inc n) = of_num n + 1" - by (induct n) (simp_all add: add_ac) - -lemma of_num_add: "of_num (m + n) = of_num m + of_num n" - by (induct n rule: num_induct) (simp_all add: add_One add_inc of_num_inc add_ac) - -lemma of_num_mult: "of_num (m * n) = of_num m * of_num n" - by (induct n rule: num_induct) (simp_all add: mult_One mult_inc of_num_add of_num_inc right_distrib) - -declare of_num.simps [simp del] - -end - -ML {* -fun mk_num k = - if k > 1 then - let - val (l, b) = Integer.div_mod k 2; - val bit = (if b = 0 then @{term Dig0} else @{term Dig1}); - in bit $ (mk_num l) end - else if k = 1 then @{term One} - else error ("mk_num " ^ string_of_int k); - -fun dest_num @{term One} = 1 - | dest_num (@{term Dig0} $ n) = 2 * dest_num n - | dest_num (@{term Dig1} $ n) = 2 * dest_num n + 1 - | dest_num t = raise TERM ("dest_num", [t]); - -fun mk_numeral phi T k = Morphism.term phi (Const (@{const_name of_num}, @{typ num} --> T)) - $ mk_num k - -fun dest_numeral phi (u $ t) = - if Term.aconv_untyped (u, Morphism.term phi (Const (@{const_name of_num}, dummyT))) - then (range_type (fastype_of u), dest_num t) - else raise TERM ("dest_numeral", [u, t]); -*} - -syntax - "_Numerals" :: "xnum_token \ 'a" ("_") - -parse_translation {* -let - fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2) - of (0, 1) => Const (@{const_name One}, dummyT) - | (n, 0) => Const (@{const_name Dig0}, dummyT) $ num_of_int n - | (n, 1) => Const (@{const_name Dig1}, dummyT) $ num_of_int n - else raise Match; - fun numeral_tr [Free (num, _)] = - let - val {leading_zeros, value, ...} = Lexicon.read_xnum num; - val _ = leading_zeros = 0 andalso value > 0 - orelse error ("Bad numeral: " ^ num); - in Const (@{const_name of_num}, @{typ num} --> dummyT) $ num_of_int value end - | numeral_tr ts = raise TERM ("numeral_tr", ts); -in [(@{syntax_const "_Numerals"}, numeral_tr)] end -*} - -typed_print_translation (advanced) {* -let - fun dig b n = b + 2 * n; - fun int_of_num' (Const (@{const_syntax Dig0}, _) $ n) = - dig 0 (int_of_num' n) - | int_of_num' (Const (@{const_syntax Dig1}, _) $ n) = - dig 1 (int_of_num' n) - | int_of_num' (Const (@{const_syntax One}, _)) = 1; - fun num_tr' ctxt T [n] = - let - val k = int_of_num' n; - val t' = Syntax.const @{syntax_const "_Numerals"} $ Syntax.free ("#" ^ string_of_int k); - in - case T of - Type (@{type_name fun}, [_, T']) => - if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t' - else Syntax.const @{syntax_const "_constrain"} $ t' $ Syntax_Phases.term_of_typ ctxt T' - | T' => if T' = dummyT then t' else raise Match - end; -in [(@{const_syntax of_num}, num_tr')] end -*} - - -subsection {* Class-specific numeral rules *} - -subsubsection {* Class @{text semiring_numeral} *} - -context semiring_numeral -begin - -abbreviation "Num1 \ of_num One" - -text {* - Alas, there is still the duplication of @{term 1}, although the - duplicated @{term 0} has disappeared. We could get rid of it by - replacing the constructor @{term 1} in @{typ num} by two - constructors @{text two} and @{text three}, resulting in a further - blow-up. But it could be worth the effort. -*} - -lemma of_num_plus_one [numeral]: - "of_num n + 1 = of_num (n + One)" - by (simp only: of_num_add of_num_One) - -lemma of_num_one_plus [numeral]: - "1 + of_num n = of_num (One + n)" - by (simp only: of_num_add of_num_One) - -lemma of_num_plus [numeral]: - "of_num m + of_num n = of_num (m + n)" - by (simp only: of_num_add) - -lemma of_num_times_one [numeral]: - "of_num n * 1 = of_num n" - by simp - -lemma of_num_one_times [numeral]: - "1 * of_num n = of_num n" - by simp - -lemma of_num_times [numeral]: - "of_num m * of_num n = of_num (m * n)" - unfolding of_num_mult .. - -end - - -subsubsection {* Structures with a zero: class @{text semiring_1} *} - -context semiring_1 -begin - -subclass semiring_numeral .. - -lemma of_nat_of_num [numeral]: "of_nat (of_num n) = of_num n" - by (induct n) - (simp_all add: semiring_numeral_class.of_num.simps of_num.simps add_ac) - -declare of_nat_1 [numeral] - -lemma Dig_plus_zero [numeral]: - "0 + 1 = 1" - "0 + of_num n = of_num n" - "1 + 0 = 1" - "of_num n + 0 = of_num n" - by simp_all - -lemma Dig_times_zero [numeral]: - "0 * 1 = 0" - "0 * of_num n = 0" - "1 * 0 = 0" - "of_num n * 0 = 0" - by simp_all - -end - -lemma nat_of_num_of_num: "nat_of_num = of_num" -proof - fix n - have "of_num n = nat_of_num n" - by (induct n) (simp_all add: of_num.simps) - then show "nat_of_num n = of_num n" by simp -qed - - -subsubsection {* Equality: class @{text semiring_char_0} *} - -context semiring_char_0 -begin - -lemma of_num_eq_iff [numeral]: "of_num m = of_num n \ m = n" - unfolding of_nat_of_num [symmetric] nat_of_num_of_num [symmetric] - of_nat_eq_iff num_eq_iff .. - -lemma of_num_eq_one_iff [numeral]: "of_num n = 1 \ n = One" - using of_num_eq_iff [of n One] by (simp add: of_num_One) - -lemma one_eq_of_num_iff [numeral]: "1 = of_num n \ One = n" - using of_num_eq_iff [of One n] by (simp add: of_num_One) - -end - - -subsubsection {* Comparisons: class @{text linordered_semidom} *} - -text {* - Perhaps the underlying structure could even - be more general than @{text linordered_semidom}. -*} - -context linordered_semidom -begin - -lemma of_num_pos [numeral]: "0 < of_num n" - by (induct n) (simp_all add: of_num.simps add_pos_pos) - -lemma of_num_not_zero [numeral]: "of_num n \ 0" - using of_num_pos [of n] by simp - -lemma of_num_less_eq_iff [numeral]: "of_num m \ of_num n \ m \ n" -proof - - have "of_nat (of_num m) \ of_nat (of_num n) \ m \ n" - unfolding less_eq_num_def nat_of_num_of_num of_nat_le_iff .. - then show ?thesis by (simp add: of_nat_of_num) -qed - -lemma of_num_less_eq_one_iff [numeral]: "of_num n \ 1 \ n \ One" - using of_num_less_eq_iff [of n One] by (simp add: of_num_One) - -lemma one_less_eq_of_num_iff [numeral]: "1 \ of_num n" - using of_num_less_eq_iff [of One n] by (simp add: of_num_One) - -lemma of_num_less_iff [numeral]: "of_num m < of_num n \ m < n" -proof - - have "of_nat (of_num m) < of_nat (of_num n) \ m < n" - unfolding less_num_def nat_of_num_of_num of_nat_less_iff .. - then show ?thesis by (simp add: of_nat_of_num) -qed - -lemma of_num_less_one_iff [numeral]: "\ of_num n < 1" - using of_num_less_iff [of n One] by (simp add: of_num_One) - -lemma one_less_of_num_iff [numeral]: "1 < of_num n \ One < n" - using of_num_less_iff [of One n] by (simp add: of_num_One) - -lemma of_num_nonneg [numeral]: "0 \ of_num n" - by (induct n) (simp_all add: of_num.simps add_nonneg_nonneg) - -lemma of_num_less_zero_iff [numeral]: "\ of_num n < 0" - by (simp add: not_less of_num_nonneg) - -lemma of_num_le_zero_iff [numeral]: "\ of_num n \ 0" - by (simp add: not_le of_num_pos) - -end - -context linordered_idom -begin - -lemma minus_of_num_less_of_num_iff: "- of_num m < of_num n" -proof - - have "- of_num m < 0" by (simp add: of_num_pos) - also have "0 < of_num n" by (simp add: of_num_pos) - finally show ?thesis . -qed - -lemma minus_of_num_not_equal_of_num: "- of_num m \ of_num n" - using minus_of_num_less_of_num_iff [of m n] by simp - -lemma minus_of_num_less_one_iff: "- of_num n < 1" - using minus_of_num_less_of_num_iff [of n One] by (simp add: of_num_One) - -lemma minus_one_less_of_num_iff: "- 1 < of_num n" - using minus_of_num_less_of_num_iff [of One n] by (simp add: of_num_One) - -lemma minus_one_less_one_iff: "- 1 < 1" - using minus_of_num_less_of_num_iff [of One One] by (simp add: of_num_One) - -lemma minus_of_num_le_of_num_iff: "- of_num m \ of_num n" - by (simp add: less_imp_le minus_of_num_less_of_num_iff) - -lemma minus_of_num_le_one_iff: "- of_num n \ 1" - by (simp add: less_imp_le minus_of_num_less_one_iff) - -lemma minus_one_le_of_num_iff: "- 1 \ of_num n" - by (simp add: less_imp_le minus_one_less_of_num_iff) - -lemma minus_one_le_one_iff: "- 1 \ 1" - by (simp add: less_imp_le minus_one_less_one_iff) - -lemma of_num_le_minus_of_num_iff: "\ of_num m \ - of_num n" - by (simp add: not_le minus_of_num_less_of_num_iff) - -lemma one_le_minus_of_num_iff: "\ 1 \ - of_num n" - by (simp add: not_le minus_of_num_less_one_iff) - -lemma of_num_le_minus_one_iff: "\ of_num n \ - 1" - by (simp add: not_le minus_one_less_of_num_iff) - -lemma one_le_minus_one_iff: "\ 1 \ - 1" - by (simp add: not_le minus_one_less_one_iff) - -lemma of_num_less_minus_of_num_iff: "\ of_num m < - of_num n" - by (simp add: not_less minus_of_num_le_of_num_iff) - -lemma one_less_minus_of_num_iff: "\ 1 < - of_num n" - by (simp add: not_less minus_of_num_le_one_iff) - -lemma of_num_less_minus_one_iff: "\ of_num n < - 1" - by (simp add: not_less minus_one_le_of_num_iff) - -lemma one_less_minus_one_iff: "\ 1 < - 1" - by (simp add: not_less minus_one_le_one_iff) - -lemmas le_signed_numeral_special [numeral] = - minus_of_num_le_of_num_iff - minus_of_num_le_one_iff - minus_one_le_of_num_iff - minus_one_le_one_iff - of_num_le_minus_of_num_iff - one_le_minus_of_num_iff - of_num_le_minus_one_iff - one_le_minus_one_iff - -lemmas less_signed_numeral_special [numeral] = - minus_of_num_less_of_num_iff - minus_of_num_not_equal_of_num - minus_of_num_less_one_iff - minus_one_less_of_num_iff - minus_one_less_one_iff - of_num_less_minus_of_num_iff - one_less_minus_of_num_iff - of_num_less_minus_one_iff - one_less_minus_one_iff - -end - -subsubsection {* Structures with subtraction: class @{text semiring_1_minus} *} - -class semiring_minus = semiring + minus + zero + - assumes minus_inverts_plus1: "a + b = c \ c - b = a" - assumes minus_minus_zero_inverts_plus1: "a + b = c \ b - c = 0 - a" -begin - -lemma minus_inverts_plus2: "a + b = c \ c - a = b" - by (simp add: add_ac minus_inverts_plus1 [of b a]) - -lemma minus_minus_zero_inverts_plus2: "a + b = c \ a - c = 0 - b" - by (simp add: add_ac minus_minus_zero_inverts_plus1 [of b a]) - -end - -class semiring_1_minus = semiring_1 + semiring_minus -begin - -lemma Dig_of_num_pos: - assumes "k + n = m" - shows "of_num m - of_num n = of_num k" - using assms by (simp add: of_num_plus minus_inverts_plus1) - -lemma Dig_of_num_zero: - shows "of_num n - of_num n = 0" - by (rule minus_inverts_plus1) simp - -lemma Dig_of_num_neg: - assumes "k + m = n" - shows "of_num m - of_num n = 0 - of_num k" - by (rule minus_minus_zero_inverts_plus1) (simp add: of_num_plus assms) - -lemmas Dig_plus_eval = - of_num_plus of_num_eq_iff Dig_plus refl [of One, THEN eqTrueI] num.inject - -simproc_setup numeral_minus ("of_num m - of_num n") = {* - let - (*TODO proper implicit use of morphism via pattern antiquotations*) - fun cdest_of_num ct = (List.last o snd o Drule.strip_comb) ct; - fun cdest_minus ct = case (rev o snd o Drule.strip_comb) ct of [n, m] => (m, n); - fun attach_num ct = (dest_num (Thm.term_of ct), ct); - fun cdifference t = (pairself (attach_num o cdest_of_num) o cdest_minus) t; - val simplify = Raw_Simplifier.rewrite false (map mk_meta_eq @{thms Dig_plus_eval}); - fun cert ck cl cj = @{thm eqTrueE} OF [@{thm meta_eq_to_obj_eq} - OF [simplify (Drule.list_comb (@{cterm "op = :: num \ _"}, - [Drule.list_comb (@{cterm "op + :: num \ _"}, [ck, cl]), cj]))]]; - in fn phi => fn _ => fn ct => case try cdifference ct - of NONE => (NONE) - | SOME ((k, ck), (l, cl)) => SOME (let val j = k - l in if j = 0 - then Raw_Simplifier.rewrite false [mk_meta_eq (Morphism.thm phi @{thm Dig_of_num_zero})] ct - else mk_meta_eq (let - val cj = Thm.cterm_of (Thm.theory_of_cterm ct) (mk_num (abs j)); - in - (if j > 0 then (Morphism.thm phi @{thm Dig_of_num_pos}) OF [cert cj cl ck] - else (Morphism.thm phi @{thm Dig_of_num_neg}) OF [cert cj ck cl]) - end) end) - end -*} - -lemma Dig_of_num_minus_zero [numeral]: - "of_num n - 0 = of_num n" - by (simp add: minus_inverts_plus1) - -lemma Dig_one_minus_zero [numeral]: - "1 - 0 = 1" - by (simp add: minus_inverts_plus1) - -lemma Dig_one_minus_one [numeral]: - "1 - 1 = 0" - by (simp add: minus_inverts_plus1) - -lemma Dig_of_num_minus_one [numeral]: - "of_num (Dig0 n) - 1 = of_num (DigM n)" - "of_num (Dig1 n) - 1 = of_num (Dig0 n)" - by (auto intro: minus_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one) - -lemma Dig_one_minus_of_num [numeral]: - "1 - of_num (Dig0 n) = 0 - of_num (DigM n)" - "1 - of_num (Dig1 n) = 0 - of_num (Dig0 n)" - by (auto intro: minus_minus_zero_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one) - -end - - -subsubsection {* Structures with negation: class @{text ring_1} *} - -context ring_1 -begin - -subclass semiring_1_minus proof -qed (simp_all add: algebra_simps) - -lemma Dig_zero_minus_of_num [numeral]: - "0 - of_num n = - of_num n" - by simp - -lemma Dig_zero_minus_one [numeral]: - "0 - 1 = - 1" - by simp - -lemma Dig_uminus_uminus [numeral]: - "- (- of_num n) = of_num n" - by simp - -lemma Dig_plus_uminus [numeral]: - "of_num m + - of_num n = of_num m - of_num n" - "- of_num m + of_num n = of_num n - of_num m" - "- of_num m + - of_num n = - (of_num m + of_num n)" - "of_num m - - of_num n = of_num m + of_num n" - "- of_num m - of_num n = - (of_num m + of_num n)" - "- of_num m - - of_num n = of_num n - of_num m" - by (simp_all add: diff_minus add_commute) - -lemma Dig_times_uminus [numeral]: - "- of_num n * of_num m = - (of_num n * of_num m)" - "of_num n * - of_num m = - (of_num n * of_num m)" - "- of_num n * - of_num m = of_num n * of_num m" - by simp_all - -lemma of_int_of_num [numeral]: "of_int (of_num n) = of_num n" -by (induct n) - (simp_all only: of_num.simps semiring_numeral_class.of_num.simps of_int_add, simp_all) - -declare of_int_1 [numeral] - -end - - -subsubsection {* Structures with exponentiation *} - -lemma of_num_square: "of_num (square x) = of_num x * of_num x" -by (induct x) - (simp_all add: of_num.simps of_num_add algebra_simps) - -lemma of_num_pow: "of_num (pow x y) = of_num x ^ of_num y" -by (induct y) - (simp_all add: of_num.simps of_num_square of_num_mult power_add) - -lemma power_of_num [numeral]: "of_num x ^ of_num y = of_num (pow x y)" - unfolding of_num_pow .. - -lemma power_zero_of_num [numeral]: - "0 ^ of_num n = (0::'a::semiring_1)" - using of_num_pos [where n=n and ?'a=nat] - by (simp add: power_0_left) - -lemma power_minus_Dig0 [numeral]: - fixes x :: "'a::ring_1" - shows "(- x) ^ of_num (Dig0 n) = x ^ of_num (Dig0 n)" - by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc) - -lemma power_minus_Dig1 [numeral]: - fixes x :: "'a::ring_1" - shows "(- x) ^ of_num (Dig1 n) = - (x ^ of_num (Dig1 n))" - by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc) - -declare power_one [numeral] - - -subsubsection {* Greetings to @{typ nat}. *} - -instance nat :: semiring_1_minus proof -qed simp_all - -lemma Suc_of_num [numeral]: "Suc (of_num n) = of_num (n + One)" - unfolding of_num_plus_one [symmetric] by simp - -lemma nat_number: - "1 = Suc 0" - "of_num One = Suc 0" - "of_num (Dig0 n) = Suc (of_num (DigM n))" - "of_num (Dig1 n) = Suc (of_num (Dig0 n))" - by (simp_all add: of_num.simps DigM_plus_one Suc_of_num) - -declare diff_0_eq_0 [numeral] - - -subsection {* Proof tools setup *} - -subsubsection {* Numeral equations as default simplification rules *} - -declare (in semiring_numeral) of_num_One [simp] -declare (in semiring_numeral) of_num_plus_one [simp] -declare (in semiring_numeral) of_num_one_plus [simp] -declare (in semiring_numeral) of_num_plus [simp] -declare (in semiring_numeral) of_num_times [simp] - -declare (in semiring_1) of_nat_of_num [simp] - -declare (in semiring_char_0) of_num_eq_iff [simp] -declare (in semiring_char_0) of_num_eq_one_iff [simp] -declare (in semiring_char_0) one_eq_of_num_iff [simp] - -declare (in linordered_semidom) of_num_pos [simp] -declare (in linordered_semidom) of_num_not_zero [simp] -declare (in linordered_semidom) of_num_less_eq_iff [simp] -declare (in linordered_semidom) of_num_less_eq_one_iff [simp] -declare (in linordered_semidom) one_less_eq_of_num_iff [simp] -declare (in linordered_semidom) of_num_less_iff [simp] -declare (in linordered_semidom) of_num_less_one_iff [simp] -declare (in linordered_semidom) one_less_of_num_iff [simp] -declare (in linordered_semidom) of_num_nonneg [simp] -declare (in linordered_semidom) of_num_less_zero_iff [simp] -declare (in linordered_semidom) of_num_le_zero_iff [simp] - -declare (in linordered_idom) le_signed_numeral_special [simp] -declare (in linordered_idom) less_signed_numeral_special [simp] - -declare (in semiring_1_minus) Dig_of_num_minus_one [simp] -declare (in semiring_1_minus) Dig_one_minus_of_num [simp] - -declare (in ring_1) Dig_plus_uminus [simp] -declare (in ring_1) of_int_of_num [simp] - -declare power_of_num [simp] -declare power_zero_of_num [simp] -declare power_minus_Dig0 [simp] -declare power_minus_Dig1 [simp] - -declare Suc_of_num [simp] - - -subsubsection {* Reorientation of equalities *} - -setup {* - Reorient_Proc.add - (fn Const(@{const_name of_num}, _) $ _ => true - | Const(@{const_name uminus}, _) $ - (Const(@{const_name of_num}, _) $ _) => true - | _ => false) -*} - -simproc_setup reorient_num ("of_num n = x" | "- of_num m = y") = Reorient_Proc.proc - - -subsubsection {* Constant folding for multiplication in semirings *} - -context semiring_numeral -begin - -lemma mult_of_num_commute: "x * of_num n = of_num n * x" -by (induct n) - (simp_all only: of_num.simps left_distrib right_distrib mult_1_left mult_1_right) - -definition - "commutes_with a b \ a * b = b * a" - -lemma commutes_with_commute: "commutes_with a b \ a * b = b * a" -unfolding commutes_with_def . - -lemma commutes_with_left_commute: "commutes_with a b \ a * (b * c) = b * (a * c)" -unfolding commutes_with_def by (simp only: mult_assoc [symmetric]) - -lemma commutes_with_numeral: "commutes_with x (of_num n)" "commutes_with (of_num n) x" -unfolding commutes_with_def by (simp_all add: mult_of_num_commute) - -lemmas mult_ac_numeral = - mult_assoc - commutes_with_commute - commutes_with_left_commute - commutes_with_numeral - -end - -ML {* -structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = -struct - val assoc_ss = HOL_ss addsimps @{thms mult_ac_numeral} - val eq_reflection = eq_reflection - fun is_numeral (Const(@{const_name of_num}, _) $ _) = true - | is_numeral _ = false; -end; - -structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); -*} - -simproc_setup semiring_assoc_fold' ("(a::'a::semiring_numeral) * b") = - {* fn phi => fn ss => fn ct => - Semiring_Times_Assoc.proc ss (Thm.term_of ct) *} - - -subsection {* Code generator setup for @{typ int} *} - -text {* Reversing standard setup *} - -lemma [code_unfold del]: "(0::int) \ Numeral0" by simp -lemma [code_unfold del]: "(1::int) \ Numeral1" by simp -declare zero_is_num_zero [code_unfold del] -declare one_is_num_one [code_unfold del] - -lemma [code, code del]: - "(1 :: int) = 1" - "(op + :: int \ int \ int) = op +" - "(uminus :: int \ int) = uminus" - "(op - :: int \ int \ int) = op -" - "(op * :: int \ int \ int) = op *" - "(HOL.equal :: int \ int \ bool) = HOL.equal" - "(op \ :: int \ int \ bool) = op \" - "(op < :: int \ int \ bool) = op <" - by rule+ - -text {* Constructors *} - -definition Pls :: "num \ int" where - [simp, code_post]: "Pls n = of_num n" - -definition Mns :: "num \ int" where - [simp, code_post]: "Mns n = - of_num n" - -code_datatype "0::int" Pls Mns - -lemmas [code_unfold] = Pls_def [symmetric] Mns_def [symmetric] - -text {* Auxiliary operations *} - -definition dup :: "int \ int" where - [simp]: "dup k = k + k" - -lemma Dig_dup [code]: - "dup 0 = 0" - "dup (Pls n) = Pls (Dig0 n)" - "dup (Mns n) = Mns (Dig0 n)" - by (simp_all add: of_num.simps) - -definition sub :: "num \ num \ int" where - [simp]: "sub m n = (of_num m - of_num n)" - -lemma Dig_sub [code]: - "sub One One = 0" - "sub (Dig0 m) One = of_num (DigM m)" - "sub (Dig1 m) One = of_num (Dig0 m)" - "sub One (Dig0 n) = - of_num (DigM n)" - "sub One (Dig1 n) = - of_num (Dig0 n)" - "sub (Dig0 m) (Dig0 n) = dup (sub m n)" - "sub (Dig1 m) (Dig1 n) = dup (sub m n)" - "sub (Dig1 m) (Dig0 n) = dup (sub m n) + 1" - "sub (Dig0 m) (Dig1 n) = dup (sub m n) - 1" - by (simp_all add: algebra_simps num_eq_iff nat_of_num_add) - -text {* Implementations *} - -lemma one_int_code [code]: - "1 = Pls One" - by simp - -lemma plus_int_code [code]: - "k + 0 = (k::int)" - "0 + l = (l::int)" - "Pls m + Pls n = Pls (m + n)" - "Pls m + Mns n = sub m n" - "Mns m + Pls n = sub n m" - "Mns m + Mns n = Mns (m + n)" - by simp_all - -lemma uminus_int_code [code]: - "uminus 0 = (0::int)" - "uminus (Pls m) = Mns m" - "uminus (Mns m) = Pls m" - by simp_all - -lemma minus_int_code [code]: - "k - 0 = (k::int)" - "0 - l = uminus (l::int)" - "Pls m - Pls n = sub m n" - "Pls m - Mns n = Pls (m + n)" - "Mns m - Pls n = Mns (m + n)" - "Mns m - Mns n = sub n m" - by simp_all - -lemma times_int_code [code]: - "k * 0 = (0::int)" - "0 * l = (0::int)" - "Pls m * Pls n = Pls (m * n)" - "Pls m * Mns n = Mns (m * n)" - "Mns m * Pls n = Mns (m * n)" - "Mns m * Mns n = Pls (m * n)" - by simp_all - -lemma eq_int_code [code]: - "HOL.equal 0 (0::int) \ True" - "HOL.equal 0 (Pls l) \ False" - "HOL.equal 0 (Mns l) \ False" - "HOL.equal (Pls k) 0 \ False" - "HOL.equal (Pls k) (Pls l) \ HOL.equal k l" - "HOL.equal (Pls k) (Mns l) \ False" - "HOL.equal (Mns k) 0 \ False" - "HOL.equal (Mns k) (Pls l) \ False" - "HOL.equal (Mns k) (Mns l) \ HOL.equal k l" - by (auto simp add: equal dest: sym) - -lemma [code nbe]: - "HOL.equal (k::int) k \ True" - by (fact equal_refl) - -lemma less_eq_int_code [code]: - "0 \ (0::int) \ True" - "0 \ Pls l \ True" - "0 \ Mns l \ False" - "Pls k \ 0 \ False" - "Pls k \ Pls l \ k \ l" - "Pls k \ Mns l \ False" - "Mns k \ 0 \ True" - "Mns k \ Pls l \ True" - "Mns k \ Mns l \ l \ k" - by simp_all - -lemma less_int_code [code]: - "0 < (0::int) \ False" - "0 < Pls l \ True" - "0 < Mns l \ False" - "Pls k < 0 \ False" - "Pls k < Pls l \ k < l" - "Pls k < Mns l \ False" - "Mns k < 0 \ True" - "Mns k < Pls l \ True" - "Mns k < Mns l \ l < k" - by simp_all - -hide_const (open) sub dup - -text {* Pretty literals *} - -ML {* -local open Code_Thingol in - -fun add_code print target = - let - fun dest_num one' dig0' dig1' thm = - let - fun dest_dig (IConst (c, _)) = if c = dig0' then 0 - else if c = dig1' then 1 - else Code_Printer.eqn_error thm "Illegal numeral expression: illegal dig" - | dest_dig _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal digit"; - fun dest_num (IConst (c, _)) = if c = one' then 1 - else Code_Printer.eqn_error thm "Illegal numeral expression: illegal leading digit" - | dest_num (t1 `$ t2) = 2 * dest_num t2 + dest_dig t1 - | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term"; - in dest_num end; - fun pretty sgn literals [one', dig0', dig1'] _ thm _ _ [(t, _)] = - (Code_Printer.str o print literals o sgn o dest_num one' dig0' dig1' thm) t - fun add_syntax (c, sgn) = Code_Target.add_const_syntax target c - (SOME (Code_Printer.complex_const_syntax - (1, ([@{const_name One}, @{const_name Dig0}, @{const_name Dig1}], - pretty sgn)))); - in - add_syntax (@{const_name Pls}, I) - #> add_syntax (@{const_name Mns}, (fn k => ~ k)) - end; - -end -*} - -hide_const (open) One Dig0 Dig1 - - -subsection {* Toy examples *} - -definition "foo \ #4 * #2 + #7 = (#8 :: nat)" -definition "bar \ #4 * #2 + #7 \ (#8 :: int) - #3" - -code_thms foo bar -export_code foo bar checking SML OCaml? Haskell? Scala? - -text {* This is an ad-hoc @{text Code_Integer} setup. *} - -setup {* - fold (add_code Code_Printer.literal_numeral) - [Code_ML.target_SML, Code_ML.target_OCaml, Code_Haskell.target, Code_Scala.target] -*} - -code_type int - (SML "IntInf.int") - (OCaml "Big'_int.big'_int") - (Haskell "Integer") - (Scala "BigInt") - (Eval "int") - -code_const "0::int" - (SML "0/ :/ IntInf.int") - (OCaml "Big'_int.zero") - (Haskell "0") - (Scala "BigInt(0)") - (Eval "0/ :/ int") - -code_const Int.pred - (SML "IntInf.- ((_), 1)") - (OCaml "Big'_int.pred'_big'_int") - (Haskell "!(_/ -/ 1)") - (Scala "!(_ -/ 1)") - (Eval "!(_/ -/ 1)") - -code_const Int.succ - (SML "IntInf.+ ((_), 1)") - (OCaml "Big'_int.succ'_big'_int") - (Haskell "!(_/ +/ 1)") - (Scala "!(_ +/ 1)") - (Eval "!(_/ +/ 1)") - -code_const "op + \ int \ int \ int" - (SML "IntInf.+ ((_), (_))") - (OCaml "Big'_int.add'_big'_int") - (Haskell infixl 6 "+") - (Scala infixl 7 "+") - (Eval infixl 8 "+") - -code_const "uminus \ int \ int" - (SML "IntInf.~") - (OCaml "Big'_int.minus'_big'_int") - (Haskell "negate") - (Scala "!(- _)") - (Eval "~/ _") - -code_const "op - \ int \ int \ int" - (SML "IntInf.- ((_), (_))") - (OCaml "Big'_int.sub'_big'_int") - (Haskell infixl 6 "-") - (Scala infixl 7 "-") - (Eval infixl 8 "-") - -code_const "op * \ int \ int \ int" - (SML "IntInf.* ((_), (_))") - (OCaml "Big'_int.mult'_big'_int") - (Haskell infixl 7 "*") - (Scala infixl 8 "*") - (Eval infixl 9 "*") - -code_const pdivmod - (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") - (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") - (Haskell "divMod/ (abs _)/ (abs _)") - (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") - (Eval "Integer.div'_mod/ (abs _)/ (abs _)") - -code_const "HOL.equal \ int \ int \ bool" - (SML "!((_ : IntInf.int) = _)") - (OCaml "Big'_int.eq'_big'_int") - (Haskell infix 4 "==") - (Scala infixl 5 "==") - (Eval infixl 6 "=") - -code_const "op \ \ int \ int \ bool" - (SML "IntInf.<= ((_), (_))") - (OCaml "Big'_int.le'_big'_int") - (Haskell infix 4 "<=") - (Scala infixl 4 "<=") - (Eval infixl 6 "<=") - -code_const "op < \ int \ int \ bool" - (SML "IntInf.< ((_), (_))") - (OCaml "Big'_int.lt'_big'_int") - (Haskell infix 4 "<") - (Scala infixl 4 "<") - (Eval infixl 6 "<") - -code_const Code_Numeral.int_of - (SML "IntInf.fromInt") - (OCaml "_") - (Haskell "toInteger") - (Scala "!_.as'_BigInt") - (Eval "_") - -export_code foo bar checking SML OCaml? Haskell? Scala? - -end diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/ex/Numeral_Representation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Numeral_Representation.thy Tue Feb 21 12:45:00 2012 +0100 @@ -0,0 +1,1116 @@ +(* Title: HOL/ex/Numeral_Representation.thy + Author: Florian Haftmann +*) + +header {* An experimental alternative numeral representation. *} + +theory Numeral_Representation +imports Main +begin + +subsection {* The @{text num} type *} + +datatype num = One | Dig0 num | Dig1 num + +text {* Increment function for type @{typ num} *} + +primrec inc :: "num \ num" where + "inc One = Dig0 One" +| "inc (Dig0 x) = Dig1 x" +| "inc (Dig1 x) = Dig0 (inc x)" + +text {* Converting between type @{typ num} and type @{typ nat} *} + +primrec nat_of_num :: "num \ nat" where + "nat_of_num One = Suc 0" +| "nat_of_num (Dig0 x) = nat_of_num x + nat_of_num x" +| "nat_of_num (Dig1 x) = Suc (nat_of_num x + nat_of_num x)" + +primrec num_of_nat :: "nat \ num" where + "num_of_nat 0 = One" +| "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" + +lemma nat_of_num_pos: "0 < nat_of_num x" + by (induct x) simp_all + +lemma nat_of_num_neq_0: "nat_of_num x \ 0" + by (induct x) simp_all + +lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" + by (induct x) simp_all + +lemma num_of_nat_double: + "0 < n \ num_of_nat (n + n) = Dig0 (num_of_nat n)" + by (induct n) simp_all + +text {* + Type @{typ num} is isomorphic to the strictly positive + natural numbers. +*} + +lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" + by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) + +lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" + by (induct n) (simp_all add: nat_of_num_inc) + +lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" +proof + assume "nat_of_num x = nat_of_num y" + then have "num_of_nat (nat_of_num x) = num_of_nat (nat_of_num y)" by simp + then show "x = y" by (simp add: nat_of_num_inverse) +qed simp + +lemma num_induct [case_names One inc]: + fixes P :: "num \ bool" + assumes One: "P One" + and inc: "\x. P x \ P (inc x)" + shows "P x" +proof - + obtain n where n: "Suc n = nat_of_num x" + by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0) + have "P (num_of_nat (Suc n))" + proof (induct n) + case 0 show ?case using One by simp + next + case (Suc n) + then have "P (inc (num_of_nat (Suc n)))" by (rule inc) + then show "P (num_of_nat (Suc (Suc n)))" by simp + qed + with n show "P x" + by (simp add: nat_of_num_inverse) +qed + +text {* + From now on, there are two possible models for @{typ num}: as + positive naturals (rule @{text "num_induct"}) and as digit + representation (rules @{text "num.induct"}, @{text "num.cases"}). + + It is not entirely clear in which context it is better to use the + one or the other, or whether the construction should be reversed. +*} + + +subsection {* Numeral operations *} + +ML {* +structure Dig_Simps = Named_Thms +( + val name = @{binding numeral} + val description = "simplification rules for numerals" +) +*} + +setup Dig_Simps.setup + +instantiation num :: "{plus,times,ord}" +begin + +definition plus_num :: "num \ num \ num" where + "m + n = num_of_nat (nat_of_num m + nat_of_num n)" + +definition times_num :: "num \ num \ num" where + "m * n = num_of_nat (nat_of_num m * nat_of_num n)" + +definition less_eq_num :: "num \ num \ bool" where + "m \ n \ nat_of_num m \ nat_of_num n" + +definition less_num :: "num \ num \ bool" where + "m < n \ nat_of_num m < nat_of_num n" + +instance .. + +end + +lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" + unfolding plus_num_def + by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) + +lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" + unfolding times_num_def + by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) + +lemma Dig_plus [numeral, simp, code]: + "One + One = Dig0 One" + "One + Dig0 m = Dig1 m" + "One + Dig1 m = Dig0 (m + One)" + "Dig0 n + One = Dig1 n" + "Dig0 n + Dig0 m = Dig0 (n + m)" + "Dig0 n + Dig1 m = Dig1 (n + m)" + "Dig1 n + One = Dig0 (n + One)" + "Dig1 n + Dig0 m = Dig1 (n + m)" + "Dig1 n + Dig1 m = Dig0 (n + m + One)" + by (simp_all add: num_eq_iff nat_of_num_add) + +lemma Dig_times [numeral, simp, code]: + "One * One = One" + "One * Dig0 n = Dig0 n" + "One * Dig1 n = Dig1 n" + "Dig0 n * One = Dig0 n" + "Dig0 n * Dig0 m = Dig0 (n * Dig0 m)" + "Dig0 n * Dig1 m = Dig0 (n * Dig1 m)" + "Dig1 n * One = Dig1 n" + "Dig1 n * Dig0 m = Dig0 (n * Dig0 m + m)" + "Dig1 n * Dig1 m = Dig1 (n * Dig1 m + m)" + by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult + left_distrib right_distrib) + +lemma less_eq_num_code [numeral, simp, code]: + "One \ n \ True" + "Dig0 m \ One \ False" + "Dig1 m \ One \ False" + "Dig0 m \ Dig0 n \ m \ n" + "Dig0 m \ Dig1 n \ m \ n" + "Dig1 m \ Dig1 n \ m \ n" + "Dig1 m \ Dig0 n \ m < n" + using nat_of_num_pos [of n] nat_of_num_pos [of m] + by (auto simp add: less_eq_num_def less_num_def) + +lemma less_num_code [numeral, simp, code]: + "m < One \ False" + "One < One \ False" + "One < Dig0 n \ True" + "One < Dig1 n \ True" + "Dig0 m < Dig0 n \ m < n" + "Dig0 m < Dig1 n \ m \ n" + "Dig1 m < Dig1 n \ m < n" + "Dig1 m < Dig0 n \ m < n" + using nat_of_num_pos [of n] nat_of_num_pos [of m] + by (auto simp add: less_eq_num_def less_num_def) + +text {* Rules using @{text One} and @{text inc} as constructors *} + +lemma add_One: "x + One = inc x" + by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) + +lemma add_inc: "x + inc y = inc (x + y)" + by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) + +lemma mult_One: "x * One = x" + by (simp add: num_eq_iff nat_of_num_mult) + +lemma mult_inc: "x * inc y = x * y + x" + by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) + +text {* A double-and-decrement function *} + +primrec DigM :: "num \ num" where + "DigM One = One" +| "DigM (Dig0 n) = Dig1 (DigM n)" +| "DigM (Dig1 n) = Dig1 (Dig0 n)" + +lemma DigM_plus_one: "DigM n + One = Dig0 n" + by (induct n) simp_all + +lemma add_One_commute: "One + n = n + One" + by (induct n) simp_all + +lemma one_plus_DigM: "One + DigM n = Dig0 n" + by (simp add: add_One_commute DigM_plus_one) + +text {* Squaring and exponentiation *} + +primrec square :: "num \ num" where + "square One = One" +| "square (Dig0 n) = Dig0 (Dig0 (square n))" +| "square (Dig1 n) = Dig1 (Dig0 (square n + n))" + +primrec pow :: "num \ num \ num" where + "pow x One = x" +| "pow x (Dig0 y) = square (pow x y)" +| "pow x (Dig1 y) = x * square (pow x y)" + + +subsection {* Binary numerals *} + +text {* + We embed binary representations into a generic algebraic + structure using @{text of_num}. +*} + +class semiring_numeral = semiring + monoid_mult +begin + +primrec of_num :: "num \ 'a" where + of_num_One [numeral]: "of_num One = 1" +| "of_num (Dig0 n) = of_num n + of_num n" +| "of_num (Dig1 n) = of_num n + of_num n + 1" + +lemma of_num_inc: "of_num (inc n) = of_num n + 1" + by (induct n) (simp_all add: add_ac) + +lemma of_num_add: "of_num (m + n) = of_num m + of_num n" + by (induct n rule: num_induct) (simp_all add: add_One add_inc of_num_inc add_ac) + +lemma of_num_mult: "of_num (m * n) = of_num m * of_num n" + by (induct n rule: num_induct) (simp_all add: mult_One mult_inc of_num_add of_num_inc right_distrib) + +declare of_num.simps [simp del] + +end + +ML {* +fun mk_num k = + if k > 1 then + let + val (l, b) = Integer.div_mod k 2; + val bit = (if b = 0 then @{term Dig0} else @{term Dig1}); + in bit $ (mk_num l) end + else if k = 1 then @{term One} + else error ("mk_num " ^ string_of_int k); + +fun dest_num @{term One} = 1 + | dest_num (@{term Dig0} $ n) = 2 * dest_num n + | dest_num (@{term Dig1} $ n) = 2 * dest_num n + 1 + | dest_num t = raise TERM ("dest_num", [t]); + +fun mk_numeral phi T k = Morphism.term phi (Const (@{const_name of_num}, @{typ num} --> T)) + $ mk_num k + +fun dest_numeral phi (u $ t) = + if Term.aconv_untyped (u, Morphism.term phi (Const (@{const_name of_num}, dummyT))) + then (range_type (fastype_of u), dest_num t) + else raise TERM ("dest_numeral", [u, t]); +*} + +syntax + "_Numerals" :: "xnum_token \ 'a" ("_") + +parse_translation {* +let + fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2) + of (0, 1) => Const (@{const_name One}, dummyT) + | (n, 0) => Const (@{const_name Dig0}, dummyT) $ num_of_int n + | (n, 1) => Const (@{const_name Dig1}, dummyT) $ num_of_int n + else raise Match; + fun numeral_tr [Free (num, _)] = + let + val {leading_zeros, value, ...} = Lexicon.read_xnum num; + val _ = leading_zeros = 0 andalso value > 0 + orelse error ("Bad numeral: " ^ num); + in Const (@{const_name of_num}, @{typ num} --> dummyT) $ num_of_int value end + | numeral_tr ts = raise TERM ("numeral_tr", ts); +in [(@{syntax_const "_Numerals"}, numeral_tr)] end +*} + +typed_print_translation (advanced) {* +let + fun dig b n = b + 2 * n; + fun int_of_num' (Const (@{const_syntax Dig0}, _) $ n) = + dig 0 (int_of_num' n) + | int_of_num' (Const (@{const_syntax Dig1}, _) $ n) = + dig 1 (int_of_num' n) + | int_of_num' (Const (@{const_syntax One}, _)) = 1; + fun num_tr' ctxt T [n] = + let + val k = int_of_num' n; + val t' = Syntax.const @{syntax_const "_Numerals"} $ Syntax.free ("#" ^ string_of_int k); + in + case T of + Type (@{type_name fun}, [_, T']) => + if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t' + else Syntax.const @{syntax_const "_constrain"} $ t' $ Syntax_Phases.term_of_typ ctxt T' + | T' => if T' = dummyT then t' else raise Match + end; +in [(@{const_syntax of_num}, num_tr')] end +*} + + +subsection {* Class-specific numeral rules *} + +subsubsection {* Class @{text semiring_numeral} *} + +context semiring_numeral +begin + +abbreviation "Num1 \ of_num One" + +text {* + Alas, there is still the duplication of @{term 1}, although the + duplicated @{term 0} has disappeared. We could get rid of it by + replacing the constructor @{term 1} in @{typ num} by two + constructors @{text two} and @{text three}, resulting in a further + blow-up. But it could be worth the effort. +*} + +lemma of_num_plus_one [numeral]: + "of_num n + 1 = of_num (n + One)" + by (simp only: of_num_add of_num_One) + +lemma of_num_one_plus [numeral]: + "1 + of_num n = of_num (One + n)" + by (simp only: of_num_add of_num_One) + +lemma of_num_plus [numeral]: + "of_num m + of_num n = of_num (m + n)" + by (simp only: of_num_add) + +lemma of_num_times_one [numeral]: + "of_num n * 1 = of_num n" + by simp + +lemma of_num_one_times [numeral]: + "1 * of_num n = of_num n" + by simp + +lemma of_num_times [numeral]: + "of_num m * of_num n = of_num (m * n)" + unfolding of_num_mult .. + +end + + +subsubsection {* Structures with a zero: class @{text semiring_1} *} + +context semiring_1 +begin + +subclass semiring_numeral .. + +lemma of_nat_of_num [numeral]: "of_nat (of_num n) = of_num n" + by (induct n) + (simp_all add: semiring_numeral_class.of_num.simps of_num.simps add_ac) + +declare of_nat_1 [numeral] + +lemma Dig_plus_zero [numeral]: + "0 + 1 = 1" + "0 + of_num n = of_num n" + "1 + 0 = 1" + "of_num n + 0 = of_num n" + by simp_all + +lemma Dig_times_zero [numeral]: + "0 * 1 = 0" + "0 * of_num n = 0" + "1 * 0 = 0" + "of_num n * 0 = 0" + by simp_all + +end + +lemma nat_of_num_of_num: "nat_of_num = of_num" +proof + fix n + have "of_num n = nat_of_num n" + by (induct n) (simp_all add: of_num.simps) + then show "nat_of_num n = of_num n" by simp +qed + + +subsubsection {* Equality: class @{text semiring_char_0} *} + +context semiring_char_0 +begin + +lemma of_num_eq_iff [numeral]: "of_num m = of_num n \ m = n" + unfolding of_nat_of_num [symmetric] nat_of_num_of_num [symmetric] + of_nat_eq_iff num_eq_iff .. + +lemma of_num_eq_one_iff [numeral]: "of_num n = 1 \ n = One" + using of_num_eq_iff [of n One] by (simp add: of_num_One) + +lemma one_eq_of_num_iff [numeral]: "1 = of_num n \ One = n" + using of_num_eq_iff [of One n] by (simp add: of_num_One) + +end + + +subsubsection {* Comparisons: class @{text linordered_semidom} *} + +text {* + Perhaps the underlying structure could even + be more general than @{text linordered_semidom}. +*} + +context linordered_semidom +begin + +lemma of_num_pos [numeral]: "0 < of_num n" + by (induct n) (simp_all add: of_num.simps add_pos_pos) + +lemma of_num_not_zero [numeral]: "of_num n \ 0" + using of_num_pos [of n] by simp + +lemma of_num_less_eq_iff [numeral]: "of_num m \ of_num n \ m \ n" +proof - + have "of_nat (of_num m) \ of_nat (of_num n) \ m \ n" + unfolding less_eq_num_def nat_of_num_of_num of_nat_le_iff .. + then show ?thesis by (simp add: of_nat_of_num) +qed + +lemma of_num_less_eq_one_iff [numeral]: "of_num n \ 1 \ n \ One" + using of_num_less_eq_iff [of n One] by (simp add: of_num_One) + +lemma one_less_eq_of_num_iff [numeral]: "1 \ of_num n" + using of_num_less_eq_iff [of One n] by (simp add: of_num_One) + +lemma of_num_less_iff [numeral]: "of_num m < of_num n \ m < n" +proof - + have "of_nat (of_num m) < of_nat (of_num n) \ m < n" + unfolding less_num_def nat_of_num_of_num of_nat_less_iff .. + then show ?thesis by (simp add: of_nat_of_num) +qed + +lemma of_num_less_one_iff [numeral]: "\ of_num n < 1" + using of_num_less_iff [of n One] by (simp add: of_num_One) + +lemma one_less_of_num_iff [numeral]: "1 < of_num n \ One < n" + using of_num_less_iff [of One n] by (simp add: of_num_One) + +lemma of_num_nonneg [numeral]: "0 \ of_num n" + by (induct n) (simp_all add: of_num.simps add_nonneg_nonneg) + +lemma of_num_less_zero_iff [numeral]: "\ of_num n < 0" + by (simp add: not_less of_num_nonneg) + +lemma of_num_le_zero_iff [numeral]: "\ of_num n \ 0" + by (simp add: not_le of_num_pos) + +end + +context linordered_idom +begin + +lemma minus_of_num_less_of_num_iff: "- of_num m < of_num n" +proof - + have "- of_num m < 0" by (simp add: of_num_pos) + also have "0 < of_num n" by (simp add: of_num_pos) + finally show ?thesis . +qed + +lemma minus_of_num_not_equal_of_num: "- of_num m \ of_num n" + using minus_of_num_less_of_num_iff [of m n] by simp + +lemma minus_of_num_less_one_iff: "- of_num n < 1" + using minus_of_num_less_of_num_iff [of n One] by (simp add: of_num_One) + +lemma minus_one_less_of_num_iff: "- 1 < of_num n" + using minus_of_num_less_of_num_iff [of One n] by (simp add: of_num_One) + +lemma minus_one_less_one_iff: "- 1 < 1" + using minus_of_num_less_of_num_iff [of One One] by (simp add: of_num_One) + +lemma minus_of_num_le_of_num_iff: "- of_num m \ of_num n" + by (simp add: less_imp_le minus_of_num_less_of_num_iff) + +lemma minus_of_num_le_one_iff: "- of_num n \ 1" + by (simp add: less_imp_le minus_of_num_less_one_iff) + +lemma minus_one_le_of_num_iff: "- 1 \ of_num n" + by (simp add: less_imp_le minus_one_less_of_num_iff) + +lemma minus_one_le_one_iff: "- 1 \ 1" + by (simp add: less_imp_le minus_one_less_one_iff) + +lemma of_num_le_minus_of_num_iff: "\ of_num m \ - of_num n" + by (simp add: not_le minus_of_num_less_of_num_iff) + +lemma one_le_minus_of_num_iff: "\ 1 \ - of_num n" + by (simp add: not_le minus_of_num_less_one_iff) + +lemma of_num_le_minus_one_iff: "\ of_num n \ - 1" + by (simp add: not_le minus_one_less_of_num_iff) + +lemma one_le_minus_one_iff: "\ 1 \ - 1" + by (simp add: not_le minus_one_less_one_iff) + +lemma of_num_less_minus_of_num_iff: "\ of_num m < - of_num n" + by (simp add: not_less minus_of_num_le_of_num_iff) + +lemma one_less_minus_of_num_iff: "\ 1 < - of_num n" + by (simp add: not_less minus_of_num_le_one_iff) + +lemma of_num_less_minus_one_iff: "\ of_num n < - 1" + by (simp add: not_less minus_one_le_of_num_iff) + +lemma one_less_minus_one_iff: "\ 1 < - 1" + by (simp add: not_less minus_one_le_one_iff) + +lemmas le_signed_numeral_special [numeral] = + minus_of_num_le_of_num_iff + minus_of_num_le_one_iff + minus_one_le_of_num_iff + minus_one_le_one_iff + of_num_le_minus_of_num_iff + one_le_minus_of_num_iff + of_num_le_minus_one_iff + one_le_minus_one_iff + +lemmas less_signed_numeral_special [numeral] = + minus_of_num_less_of_num_iff + minus_of_num_not_equal_of_num + minus_of_num_less_one_iff + minus_one_less_of_num_iff + minus_one_less_one_iff + of_num_less_minus_of_num_iff + one_less_minus_of_num_iff + of_num_less_minus_one_iff + one_less_minus_one_iff + +end + +subsubsection {* Structures with subtraction: class @{text semiring_1_minus} *} + +class semiring_minus = semiring + minus + zero + + assumes minus_inverts_plus1: "a + b = c \ c - b = a" + assumes minus_minus_zero_inverts_plus1: "a + b = c \ b - c = 0 - a" +begin + +lemma minus_inverts_plus2: "a + b = c \ c - a = b" + by (simp add: add_ac minus_inverts_plus1 [of b a]) + +lemma minus_minus_zero_inverts_plus2: "a + b = c \ a - c = 0 - b" + by (simp add: add_ac minus_minus_zero_inverts_plus1 [of b a]) + +end + +class semiring_1_minus = semiring_1 + semiring_minus +begin + +lemma Dig_of_num_pos: + assumes "k + n = m" + shows "of_num m - of_num n = of_num k" + using assms by (simp add: of_num_plus minus_inverts_plus1) + +lemma Dig_of_num_zero: + shows "of_num n - of_num n = 0" + by (rule minus_inverts_plus1) simp + +lemma Dig_of_num_neg: + assumes "k + m = n" + shows "of_num m - of_num n = 0 - of_num k" + by (rule minus_minus_zero_inverts_plus1) (simp add: of_num_plus assms) + +lemmas Dig_plus_eval = + of_num_plus of_num_eq_iff Dig_plus refl [of One, THEN eqTrueI] num.inject + +simproc_setup numeral_minus ("of_num m - of_num n") = {* + let + (*TODO proper implicit use of morphism via pattern antiquotations*) + fun cdest_of_num ct = (List.last o snd o Drule.strip_comb) ct; + fun cdest_minus ct = case (rev o snd o Drule.strip_comb) ct of [n, m] => (m, n); + fun attach_num ct = (dest_num (Thm.term_of ct), ct); + fun cdifference t = (pairself (attach_num o cdest_of_num) o cdest_minus) t; + val simplify = Raw_Simplifier.rewrite false (map mk_meta_eq @{thms Dig_plus_eval}); + fun cert ck cl cj = @{thm eqTrueE} OF [@{thm meta_eq_to_obj_eq} + OF [simplify (Drule.list_comb (@{cterm "op = :: num \ _"}, + [Drule.list_comb (@{cterm "op + :: num \ _"}, [ck, cl]), cj]))]]; + in fn phi => fn _ => fn ct => case try cdifference ct + of NONE => (NONE) + | SOME ((k, ck), (l, cl)) => SOME (let val j = k - l in if j = 0 + then Raw_Simplifier.rewrite false [mk_meta_eq (Morphism.thm phi @{thm Dig_of_num_zero})] ct + else mk_meta_eq (let + val cj = Thm.cterm_of (Thm.theory_of_cterm ct) (mk_num (abs j)); + in + (if j > 0 then (Morphism.thm phi @{thm Dig_of_num_pos}) OF [cert cj cl ck] + else (Morphism.thm phi @{thm Dig_of_num_neg}) OF [cert cj ck cl]) + end) end) + end +*} + +lemma Dig_of_num_minus_zero [numeral]: + "of_num n - 0 = of_num n" + by (simp add: minus_inverts_plus1) + +lemma Dig_one_minus_zero [numeral]: + "1 - 0 = 1" + by (simp add: minus_inverts_plus1) + +lemma Dig_one_minus_one [numeral]: + "1 - 1 = 0" + by (simp add: minus_inverts_plus1) + +lemma Dig_of_num_minus_one [numeral]: + "of_num (Dig0 n) - 1 = of_num (DigM n)" + "of_num (Dig1 n) - 1 = of_num (Dig0 n)" + by (auto intro: minus_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one) + +lemma Dig_one_minus_of_num [numeral]: + "1 - of_num (Dig0 n) = 0 - of_num (DigM n)" + "1 - of_num (Dig1 n) = 0 - of_num (Dig0 n)" + by (auto intro: minus_minus_zero_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one) + +end + + +subsubsection {* Structures with negation: class @{text ring_1} *} + +context ring_1 +begin + +subclass semiring_1_minus proof +qed (simp_all add: algebra_simps) + +lemma Dig_zero_minus_of_num [numeral]: + "0 - of_num n = - of_num n" + by simp + +lemma Dig_zero_minus_one [numeral]: + "0 - 1 = - 1" + by simp + +lemma Dig_uminus_uminus [numeral]: + "- (- of_num n) = of_num n" + by simp + +lemma Dig_plus_uminus [numeral]: + "of_num m + - of_num n = of_num m - of_num n" + "- of_num m + of_num n = of_num n - of_num m" + "- of_num m + - of_num n = - (of_num m + of_num n)" + "of_num m - - of_num n = of_num m + of_num n" + "- of_num m - of_num n = - (of_num m + of_num n)" + "- of_num m - - of_num n = of_num n - of_num m" + by (simp_all add: diff_minus add_commute) + +lemma Dig_times_uminus [numeral]: + "- of_num n * of_num m = - (of_num n * of_num m)" + "of_num n * - of_num m = - (of_num n * of_num m)" + "- of_num n * - of_num m = of_num n * of_num m" + by simp_all + +lemma of_int_of_num [numeral]: "of_int (of_num n) = of_num n" +by (induct n) + (simp_all only: of_num.simps semiring_numeral_class.of_num.simps of_int_add, simp_all) + +declare of_int_1 [numeral] + +end + + +subsubsection {* Structures with exponentiation *} + +lemma of_num_square: "of_num (square x) = of_num x * of_num x" +by (induct x) + (simp_all add: of_num.simps of_num_add algebra_simps) + +lemma of_num_pow: "of_num (pow x y) = of_num x ^ of_num y" +by (induct y) + (simp_all add: of_num.simps of_num_square of_num_mult power_add) + +lemma power_of_num [numeral]: "of_num x ^ of_num y = of_num (pow x y)" + unfolding of_num_pow .. + +lemma power_zero_of_num [numeral]: + "0 ^ of_num n = (0::'a::semiring_1)" + using of_num_pos [where n=n and ?'a=nat] + by (simp add: power_0_left) + +lemma power_minus_Dig0 [numeral]: + fixes x :: "'a::ring_1" + shows "(- x) ^ of_num (Dig0 n) = x ^ of_num (Dig0 n)" + by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc) + +lemma power_minus_Dig1 [numeral]: + fixes x :: "'a::ring_1" + shows "(- x) ^ of_num (Dig1 n) = - (x ^ of_num (Dig1 n))" + by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc) + +declare power_one [numeral] + + +subsubsection {* Greetings to @{typ nat}. *} + +instance nat :: semiring_1_minus proof +qed simp_all + +lemma Suc_of_num [numeral]: "Suc (of_num n) = of_num (n + One)" + unfolding of_num_plus_one [symmetric] by simp + +lemma nat_number: + "1 = Suc 0" + "of_num One = Suc 0" + "of_num (Dig0 n) = Suc (of_num (DigM n))" + "of_num (Dig1 n) = Suc (of_num (Dig0 n))" + by (simp_all add: of_num.simps DigM_plus_one Suc_of_num) + +declare diff_0_eq_0 [numeral] + + +subsection {* Proof tools setup *} + +subsubsection {* Numeral equations as default simplification rules *} + +declare (in semiring_numeral) of_num_One [simp] +declare (in semiring_numeral) of_num_plus_one [simp] +declare (in semiring_numeral) of_num_one_plus [simp] +declare (in semiring_numeral) of_num_plus [simp] +declare (in semiring_numeral) of_num_times [simp] + +declare (in semiring_1) of_nat_of_num [simp] + +declare (in semiring_char_0) of_num_eq_iff [simp] +declare (in semiring_char_0) of_num_eq_one_iff [simp] +declare (in semiring_char_0) one_eq_of_num_iff [simp] + +declare (in linordered_semidom) of_num_pos [simp] +declare (in linordered_semidom) of_num_not_zero [simp] +declare (in linordered_semidom) of_num_less_eq_iff [simp] +declare (in linordered_semidom) of_num_less_eq_one_iff [simp] +declare (in linordered_semidom) one_less_eq_of_num_iff [simp] +declare (in linordered_semidom) of_num_less_iff [simp] +declare (in linordered_semidom) of_num_less_one_iff [simp] +declare (in linordered_semidom) one_less_of_num_iff [simp] +declare (in linordered_semidom) of_num_nonneg [simp] +declare (in linordered_semidom) of_num_less_zero_iff [simp] +declare (in linordered_semidom) of_num_le_zero_iff [simp] + +declare (in linordered_idom) le_signed_numeral_special [simp] +declare (in linordered_idom) less_signed_numeral_special [simp] + +declare (in semiring_1_minus) Dig_of_num_minus_one [simp] +declare (in semiring_1_minus) Dig_one_minus_of_num [simp] + +declare (in ring_1) Dig_plus_uminus [simp] +declare (in ring_1) of_int_of_num [simp] + +declare power_of_num [simp] +declare power_zero_of_num [simp] +declare power_minus_Dig0 [simp] +declare power_minus_Dig1 [simp] + +declare Suc_of_num [simp] + + +subsubsection {* Reorientation of equalities *} + +setup {* + Reorient_Proc.add + (fn Const(@{const_name of_num}, _) $ _ => true + | Const(@{const_name uminus}, _) $ + (Const(@{const_name of_num}, _) $ _) => true + | _ => false) +*} + +simproc_setup reorient_num ("of_num n = x" | "- of_num m = y") = Reorient_Proc.proc + + +subsubsection {* Constant folding for multiplication in semirings *} + +context semiring_numeral +begin + +lemma mult_of_num_commute: "x * of_num n = of_num n * x" +by (induct n) + (simp_all only: of_num.simps left_distrib right_distrib mult_1_left mult_1_right) + +definition + "commutes_with a b \ a * b = b * a" + +lemma commutes_with_commute: "commutes_with a b \ a * b = b * a" +unfolding commutes_with_def . + +lemma commutes_with_left_commute: "commutes_with a b \ a * (b * c) = b * (a * c)" +unfolding commutes_with_def by (simp only: mult_assoc [symmetric]) + +lemma commutes_with_numeral: "commutes_with x (of_num n)" "commutes_with (of_num n) x" +unfolding commutes_with_def by (simp_all add: mult_of_num_commute) + +lemmas mult_ac_numeral = + mult_assoc + commutes_with_commute + commutes_with_left_commute + commutes_with_numeral + +end + +ML {* +structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = +struct + val assoc_ss = HOL_ss addsimps @{thms mult_ac_numeral} + val eq_reflection = eq_reflection + fun is_numeral (Const(@{const_name of_num}, _) $ _) = true + | is_numeral _ = false; +end; + +structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); +*} + +simproc_setup semiring_assoc_fold' ("(a::'a::semiring_numeral) * b") = + {* fn phi => fn ss => fn ct => + Semiring_Times_Assoc.proc ss (Thm.term_of ct) *} + + +subsection {* Code generator setup for @{typ int} *} + +text {* Reversing standard setup *} + +lemma [code_unfold del]: "(0::int) \ Numeral0" by simp +lemma [code_unfold del]: "(1::int) \ Numeral1" by simp +declare zero_is_num_zero [code_unfold del] +declare one_is_num_one [code_unfold del] + +lemma [code, code del]: + "(1 :: int) = 1" + "(op + :: int \ int \ int) = op +" + "(uminus :: int \ int) = uminus" + "(op - :: int \ int \ int) = op -" + "(op * :: int \ int \ int) = op *" + "(HOL.equal :: int \ int \ bool) = HOL.equal" + "(op \ :: int \ int \ bool) = op \" + "(op < :: int \ int \ bool) = op <" + by rule+ + +text {* Constructors *} + +definition Pls :: "num \ int" where + [simp, code_post]: "Pls n = of_num n" + +definition Mns :: "num \ int" where + [simp, code_post]: "Mns n = - of_num n" + +code_datatype "0::int" Pls Mns + +lemmas [code_unfold] = Pls_def [symmetric] Mns_def [symmetric] + +text {* Auxiliary operations *} + +definition dup :: "int \ int" where + [simp]: "dup k = k + k" + +lemma Dig_dup [code]: + "dup 0 = 0" + "dup (Pls n) = Pls (Dig0 n)" + "dup (Mns n) = Mns (Dig0 n)" + by (simp_all add: of_num.simps) + +definition sub :: "num \ num \ int" where + [simp]: "sub m n = (of_num m - of_num n)" + +lemma Dig_sub [code]: + "sub One One = 0" + "sub (Dig0 m) One = of_num (DigM m)" + "sub (Dig1 m) One = of_num (Dig0 m)" + "sub One (Dig0 n) = - of_num (DigM n)" + "sub One (Dig1 n) = - of_num (Dig0 n)" + "sub (Dig0 m) (Dig0 n) = dup (sub m n)" + "sub (Dig1 m) (Dig1 n) = dup (sub m n)" + "sub (Dig1 m) (Dig0 n) = dup (sub m n) + 1" + "sub (Dig0 m) (Dig1 n) = dup (sub m n) - 1" + by (simp_all add: algebra_simps num_eq_iff nat_of_num_add) + +text {* Implementations *} + +lemma one_int_code [code]: + "1 = Pls One" + by simp + +lemma plus_int_code [code]: + "k + 0 = (k::int)" + "0 + l = (l::int)" + "Pls m + Pls n = Pls (m + n)" + "Pls m + Mns n = sub m n" + "Mns m + Pls n = sub n m" + "Mns m + Mns n = Mns (m + n)" + by simp_all + +lemma uminus_int_code [code]: + "uminus 0 = (0::int)" + "uminus (Pls m) = Mns m" + "uminus (Mns m) = Pls m" + by simp_all + +lemma minus_int_code [code]: + "k - 0 = (k::int)" + "0 - l = uminus (l::int)" + "Pls m - Pls n = sub m n" + "Pls m - Mns n = Pls (m + n)" + "Mns m - Pls n = Mns (m + n)" + "Mns m - Mns n = sub n m" + by simp_all + +lemma times_int_code [code]: + "k * 0 = (0::int)" + "0 * l = (0::int)" + "Pls m * Pls n = Pls (m * n)" + "Pls m * Mns n = Mns (m * n)" + "Mns m * Pls n = Mns (m * n)" + "Mns m * Mns n = Pls (m * n)" + by simp_all + +lemma eq_int_code [code]: + "HOL.equal 0 (0::int) \ True" + "HOL.equal 0 (Pls l) \ False" + "HOL.equal 0 (Mns l) \ False" + "HOL.equal (Pls k) 0 \ False" + "HOL.equal (Pls k) (Pls l) \ HOL.equal k l" + "HOL.equal (Pls k) (Mns l) \ False" + "HOL.equal (Mns k) 0 \ False" + "HOL.equal (Mns k) (Pls l) \ False" + "HOL.equal (Mns k) (Mns l) \ HOL.equal k l" + by (auto simp add: equal dest: sym) + +lemma [code nbe]: + "HOL.equal (k::int) k \ True" + by (fact equal_refl) + +lemma less_eq_int_code [code]: + "0 \ (0::int) \ True" + "0 \ Pls l \ True" + "0 \ Mns l \ False" + "Pls k \ 0 \ False" + "Pls k \ Pls l \ k \ l" + "Pls k \ Mns l \ False" + "Mns k \ 0 \ True" + "Mns k \ Pls l \ True" + "Mns k \ Mns l \ l \ k" + by simp_all + +lemma less_int_code [code]: + "0 < (0::int) \ False" + "0 < Pls l \ True" + "0 < Mns l \ False" + "Pls k < 0 \ False" + "Pls k < Pls l \ k < l" + "Pls k < Mns l \ False" + "Mns k < 0 \ True" + "Mns k < Pls l \ True" + "Mns k < Mns l \ l < k" + by simp_all + +hide_const (open) sub dup + +text {* Pretty literals *} + +ML {* +local open Code_Thingol in + +fun add_code print target = + let + fun dest_num one' dig0' dig1' thm = + let + fun dest_dig (IConst (c, _)) = if c = dig0' then 0 + else if c = dig1' then 1 + else Code_Printer.eqn_error thm "Illegal numeral expression: illegal dig" + | dest_dig _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal digit"; + fun dest_num (IConst (c, _)) = if c = one' then 1 + else Code_Printer.eqn_error thm "Illegal numeral expression: illegal leading digit" + | dest_num (t1 `$ t2) = 2 * dest_num t2 + dest_dig t1 + | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term"; + in dest_num end; + fun pretty sgn literals [one', dig0', dig1'] _ thm _ _ [(t, _)] = + (Code_Printer.str o print literals o sgn o dest_num one' dig0' dig1' thm) t + fun add_syntax (c, sgn) = Code_Target.add_const_syntax target c + (SOME (Code_Printer.complex_const_syntax + (1, ([@{const_name One}, @{const_name Dig0}, @{const_name Dig1}], + pretty sgn)))); + in + add_syntax (@{const_name Pls}, I) + #> add_syntax (@{const_name Mns}, (fn k => ~ k)) + end; + +end +*} + +hide_const (open) One Dig0 Dig1 + + +subsection {* Toy examples *} + +definition "foo \ #4 * #2 + #7 = (#8 :: nat)" +definition "bar \ #4 * #2 + #7 \ (#8 :: int) - #3" + +code_thms foo bar +export_code foo bar checking SML OCaml? Haskell? Scala? + +text {* This is an ad-hoc @{text Code_Integer} setup. *} + +setup {* + fold (add_code Code_Printer.literal_numeral) + [Code_ML.target_SML, Code_ML.target_OCaml, Code_Haskell.target, Code_Scala.target] +*} + +code_type int + (SML "IntInf.int") + (OCaml "Big'_int.big'_int") + (Haskell "Integer") + (Scala "BigInt") + (Eval "int") + +code_const "0::int" + (SML "0/ :/ IntInf.int") + (OCaml "Big'_int.zero") + (Haskell "0") + (Scala "BigInt(0)") + (Eval "0/ :/ int") + +code_const Int.pred + (SML "IntInf.- ((_), 1)") + (OCaml "Big'_int.pred'_big'_int") + (Haskell "!(_/ -/ 1)") + (Scala "!(_ -/ 1)") + (Eval "!(_/ -/ 1)") + +code_const Int.succ + (SML "IntInf.+ ((_), 1)") + (OCaml "Big'_int.succ'_big'_int") + (Haskell "!(_/ +/ 1)") + (Scala "!(_ +/ 1)") + (Eval "!(_/ +/ 1)") + +code_const "op + \ int \ int \ int" + (SML "IntInf.+ ((_), (_))") + (OCaml "Big'_int.add'_big'_int") + (Haskell infixl 6 "+") + (Scala infixl 7 "+") + (Eval infixl 8 "+") + +code_const "uminus \ int \ int" + (SML "IntInf.~") + (OCaml "Big'_int.minus'_big'_int") + (Haskell "negate") + (Scala "!(- _)") + (Eval "~/ _") + +code_const "op - \ int \ int \ int" + (SML "IntInf.- ((_), (_))") + (OCaml "Big'_int.sub'_big'_int") + (Haskell infixl 6 "-") + (Scala infixl 7 "-") + (Eval infixl 8 "-") + +code_const "op * \ int \ int \ int" + (SML "IntInf.* ((_), (_))") + (OCaml "Big'_int.mult'_big'_int") + (Haskell infixl 7 "*") + (Scala infixl 8 "*") + (Eval infixl 9 "*") + +code_const pdivmod + (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") + (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") + (Haskell "divMod/ (abs _)/ (abs _)") + (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") + (Eval "Integer.div'_mod/ (abs _)/ (abs _)") + +code_const "HOL.equal \ int \ int \ bool" + (SML "!((_ : IntInf.int) = _)") + (OCaml "Big'_int.eq'_big'_int") + (Haskell infix 4 "==") + (Scala infixl 5 "==") + (Eval infixl 6 "=") + +code_const "op \ \ int \ int \ bool" + (SML "IntInf.<= ((_), (_))") + (OCaml "Big'_int.le'_big'_int") + (Haskell infix 4 "<=") + (Scala infixl 4 "<=") + (Eval infixl 6 "<=") + +code_const "op < \ int \ int \ bool" + (SML "IntInf.< ((_), (_))") + (OCaml "Big'_int.lt'_big'_int") + (Haskell infix 4 "<") + (Scala infixl 4 "<") + (Eval infixl 6 "<") + +code_const Code_Numeral.int_of + (SML "IntInf.fromInt") + (OCaml "_") + (Haskell "toInteger") + (Scala "!_.as'_BigInt") + (Eval "_") + +export_code foo bar checking SML OCaml? Haskell? Scala? + +end diff -r 0133d31f9ab4 -r 26977429b784 src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Mon Feb 20 22:35:32 2012 +0100 +++ b/src/HOL/ex/ROOT.ML Tue Feb 21 12:45:00 2012 +0100 @@ -17,7 +17,7 @@ use_thys [ "Iff_Oracle", "Coercion_Examples", - "Numeral", + "Numeral_Representation", "Higher_Order_Logic", "Abstract_NAT", "Guess",