diff -r 000000000000 -r 7949f97df77a arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/arith.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,229 @@ +(* Title: HOL/arith.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For HOL/arith.thy. + +Proofs about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +open Arith; + +(*** Basic rewrite and congruence rules for the arithmetic operators ***) + +val [add_0,add_Suc] = nat_recs add_def; +val [mult_0,mult_Suc] = nat_recs mult_def; + +(** Difference **) + +val diff_0 = diff_def RS def_nat_rec_0; + +val diff_0_eq_0 = prove_goalw Arith.thy [diff_def] "0 - n = 0" + (fn _ => [nat_ind_tac "n" 1, ALLGOALS(asm_simp_tac nat_ss)]); + +(*Must simplify BEFORE the induction!! (Else we get a critical pair) + Suc(m) - Suc(n) rewrites to pred(Suc(m) - n) *) +val diff_Suc_Suc = prove_goalw Arith.thy [diff_def] "Suc(m) - Suc(n) = m - n" + (fn _ => + [simp_tac nat_ss 1, nat_ind_tac "n" 1, ALLGOALS(asm_simp_tac nat_ss)]); + +(*** Simplification over add, mult, diff ***) + +val arith_simps = [add_0, add_Suc, + mult_0, mult_Suc, + diff_0, diff_0_eq_0, diff_Suc_Suc]; + +val arith_ss = nat_ss addsimps arith_simps; + +(**** Inductive properties of the operators ****) + +(*** Addition ***) + +val add_0_right = prove_goal Arith.thy "m + 0 = m" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +val add_Suc_right = prove_goal Arith.thy "m + Suc(n) = Suc(m+n)" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +val arith_ss = arith_ss addsimps [add_0_right,add_Suc_right]; + +(*Associative law for addition*) +val add_assoc = prove_goal Arith.thy "(m + n) + k = m + (n + k)::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*Commutative law for addition. Must simplify after first induction! + Orientation of rewrites is delicate*) +val add_commute = prove_goal Arith.thy "m + n = n + m::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*** Multiplication ***) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy "m * 0 = 0" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*right Sucessor law for multiplication*) +val mult_Suc_right = prove_goal Arith.thy + "m * Suc(n) = m + (m * n)" + (fn _ => + [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_assoc RS sym])), + (*The final goal requires the commutative law for addition*) + stac add_commute 1, rtac refl 1]); + +val arith_ss = arith_ss addsimps [mult_0_right,mult_Suc_right]; + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy "m * n = n * m::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS (asm_simp_tac arith_ss)]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy "(m + n)*k = (m*k) + (n*k)::nat" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_assoc]))]); + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy "(m * n) * k = m * (n * k)::nat" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_mult_distrib]))]); + + +(*** Difference ***) + +val diff_self_eq_0 = prove_goal Arith.thy "m - m = 0" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *) +val [prem] = goal Arith.thy "[| ~ m n + (m-n) = m::nat"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS(asm_simp_tac arith_ss)); +val add_diff_inverse = result(); + + +(*** Remainder ***) + +goal Arith.thy "m - n < Suc(m)"; +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (etac less_SucE 3); +by (ALLGOALS(asm_simp_tac arith_ss)); +val diff_less_Suc = result(); + +(*In ordinary notation: if 0 m - n < m"; +by (subgoal_tac "0 ~ m m - n < m" 1); +by (fast_tac HOL_cs 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS(asm_simp_tac(arith_ss addsimps [diff_less_Suc]))); +val div_termination = result(); + +val wf_less_trans = wf_pred_nat RS wf_trancl RSN (2, def_wfrec RS trans); + +goalw Nat.thy [less_def] " : pred_nat^+ = (m m mod n = m"; +by (rtac (mod_def RS wf_less_trans) 1); +by(asm_simp_tac HOL_ss 1); +val mod_less = result(); + +goal Arith.thy "!!m. [| 0 m mod n = (m-n) mod n"; +by (rtac (mod_def RS wf_less_trans) 1); +by(asm_simp_tac (nat_ss addsimps [div_termination, cut_apply, less_eq]) 1); +val mod_geq = result(); + + +(*** Quotient ***) + +goal Arith.thy "!!m. m m div n = 0"; +by (rtac (div_def RS wf_less_trans) 1); +by(asm_simp_tac nat_ss 1); +val div_less = result(); + +goal Arith.thy "!!M. [| 0 m div n = Suc((m-n) div n)"; +by (rtac (div_def RS wf_less_trans) 1); +by(asm_simp_tac (nat_ss addsimps [div_termination, cut_apply, less_eq]) 1); +val div_geq = result(); + +(*Main Result about quotient and remainder.*) +goal Arith.thy "!!m. 0 (m div n)*n + m mod n = m"; +by (res_inst_tac [("n","m")] less_induct 1); +by (rename_tac "k" 1); (*Variable name used in line below*) +by (res_inst_tac [("Q","k m-n = 0"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +val less_imp_diff_is_0 = result(); + +val prems = goal Arith.thy "m-n = 0 --> n-m = 0 --> m=n"; +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (REPEAT(simp_tac arith_ss 1 THEN TRY(atac 1))); +val diffs0_imp_equal_lemma = result(); + +(* [| m-n = 0; n-m = 0 |] ==> m=n *) +val diffs0_imp_equal = standard (diffs0_imp_equal_lemma RS mp RS mp); + +val [prem] = goal Arith.thy "m 0 Suc(m)-n = Suc(m-n)"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS(asm_simp_tac arith_ss)); +val Suc_diff_n = result(); + +goal Arith.thy "Suc(m)-n = if(m (!n. P(Suc(n))--> P(n)) --> P(k-i)"; +by (res_inst_tac [("m","k"),("n","i")] diff_induct 1); +by (ALLGOALS (strip_tac THEN' simp_tac arith_ss THEN' fast_tac HOL_cs)); +val zero_induct_lemma = result(); + +val prems = goal Arith.thy "[| P(k); !!n. P(Suc(n)) ==> P(n) |] ==> P(0)"; +by (rtac (diff_self_eq_0 RS subst) 1); +by (rtac (zero_induct_lemma RS mp RS mp) 1); +by (REPEAT (ares_tac ([impI,allI]@prems) 1)); +val zero_induct = result(); + +(*13 July 1992: loaded in 105.7s*) + +(**** Additional theorems about "less than" ****) + +goal Arith.thy "n <= (m + n)::nat"; +by (nat_ind_tac "m" 1); +by (ALLGOALS(simp_tac (arith_ss addsimps [le_refl]))); +by (etac le_trans 1); +by (rtac (lessI RS less_imp_le) 1); +val le_add2 = result(); + +goal Arith.thy "m <= (m + n)::nat"; +by (stac add_commute 1); +by (rtac le_add2 1); +val le_add1 = result(); + +val less_add_Suc1 = standard (lessI RS (le_add1 RS le_less_trans)); +val less_add_Suc2 = standard (lessI RS (le_add2 RS le_less_trans));