Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
split_weak_cong, nat_case_weak_cong, nat_rec_weak_cong. Used in llist.ML
to make simplifications faster.
HOL/gfp: re-ordered premises to put mono(f) early (first or right after
A==gfp(f) in the def_ rules). Renamed some variables in rules, A to X and
h to A. Renamed coinduct to weak_coinduct, coinduct2 to coinduct.
Strengthened coinduct as suggested by j. Frost, to have the premise X <= f(X
Un gfp(f)).
HOL/llist: used stronger coinduct rule to strengthen LList_coinduct,
LList_equalityI, llist_equalityI, llist_fun_equalityI and to delete the "2"
form of those rules. Proved List_Fun_LList_I, LListD_Fun_diag_I and
llistD_Fun_range_I to help use the new coinduction rules; most proofs
involving them required small changes. Proved prod_fun_range_eq_diag as
lemma for llist_equalityI.
(* Title: HOL/arith.thy
ID: $Id$
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1993 University of Cambridge
Arithmetic operators and their definitions
*)
Arith = Nat +
arities nat::plus
nat::minus
nat::times
consts
div,mod :: "[nat,nat]=>nat" (infixl 70)
rules
add_def "m+n == nat_rec(m, n, %u v.Suc(v))"
diff_def "m-n == nat_rec(n, m, %u v. nat_rec(v, 0, %x y.x))"
mult_def "m*n == nat_rec(m, 0, %u v. n + v)"
mod_def "m mod n == wfrec(trancl(pred_nat), m, %j f. if(j<n, j, f(j-n)))"
div_def "m div n == wfrec(trancl(pred_nat), m, %j f. if(j<n, 0, Suc(f(j-n))))"
end
(*"Difference" is subtraction of natural numbers.
There are no negative numbers; we have
m - n = 0 iff m<=n and m - n = Suc(k) iff m>n.
Also, nat_rec(m, 0, %z w.z) is pred(m). *)