Wed 22nd Jul 2015 Library/wo.ml Added the basic fact that a finite toset is a woset: WOSET_FINITE_TOSET = |- !l. toset l /\ FINITE (fl l) ==> woset l Mon 20th Jul 2015 sets.ml Added a series of theorems about injectivity and surjectivity of the "preimage" construct. These are a natural dual to existing "IMAGE" versions, and essentially show that the preimage map is injective/surjective when the function itself is surjective/injective. INJECTIVE_ON_PREIMAGE = |- !f:A->B s u. (!t t'. t SUBSET u /\ t' SUBSET u /\ {x | x IN s /\ f x IN t} = {x | x IN s /\ f x IN t'} ==> t = t') <=> u SUBSET IMAGE f s SURJECTIVE_ON_PREIMAGE = |- !f s u. (!k. k SUBSET s ==> ?t. t SUBSET u /\ {x | x IN s /\ f x IN t} = k) <=> IMAGE f s SUBSET u /\ (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) INJECTIVE_PREIMAGE = |- !f. (!t t'. {x | f x IN t} = {x | f x IN t'} ==> t = t') <=> IMAGE f UNIV = UNIV SURJECTIVE_PREIMAGE = |- !f. (!k. ?t. {x | f x IN t} = k) <=> (!x y. f x = f y ==> x = y) Sun 19th Jul 2015 Library/card.ml Renamed CARD_ADD_ABSORB to CARD_ADD_ABSORB_LEFT and added a few more basic theorems: CARD_ADD_ABSORB_RIGHT = |- !s t. INFINITE s /\ t <=_c s ==> s +_c t =_c s CARD_DIFF_ABSORB = |- !s t. INFINITE s /\ t <_c s ==> s DIFF t =_c s CARD_UNION_ABSORB_LEFT = |- !s t. INFINITE t /\ s <=_c t ==> s UNION t =_c t CARD_UNION_ABSORB_RIGHT = |- !s t. INFINITE s /\ t <=_c s ==> s UNION t =_c s Fri 17th Jul 2015 Library/floor.ml Added one more theorem about the rationals: INFINITE_RATIONAL_IN_RANGE = |- !a b. a < b ==> INFINITE {q | rational q /\ a < q /\ q < b} Mon 13th Jul 2015 Library/prime.ml Added one more natural theorem about primality on N, as well as slightly reorganizing a couple of existing proofs: PRIME_IRREDUCIBLE = |- !p. prime p <=> p > 1 /\ (!a b. p divides a * b ==> p divides a \/ p divides b) Fri 10th Jul 2015 Library/floor.ml Added a couple more rational approximation theorems: RATIONAL_APPROXIMATION_ABOVE = |- !x e. &0 < e ==> (?q. rational q /\ x < q /\ q < x + e) RATIONAL_APPROXIMATION_BELOW = |- !x e. &0 < e ==> (?q. rational q /\ x - e < q /\ q < x) Fri 10th Jul 2015 Library/wo.ml, Library/card.ml Added a few new theorems about ordered sets, ordinals and cardinals, including slightly reshuffling things between these two files. Tukey's lemma is the only real "named" theorem. Several of these slightly fill out the rather sparse theory of ordinals (for example showing they are themselves wellordered by "initial segment of"), which however remains pretty sparse. INSEG_ANTISYM = |- !l m. l inseg m /\ m inseg l ==> l = m INSEG_ORDINAL = |- !l m. m inseg l /\ ordinal l ==> ordinal m INSEG_REFL = |- !l. l inseg l INSEG_TRANS = |- !l m n. l inseg m /\ m inseg n ==> l inseg n LE_C_IMAGE_SUBSET = |- !s t. s <=_c t <=> (?f. s SUBSET IMAGE f t) ORDINAL_FL_SUBSET = |- !l m. ordinal l /\ ordinal m /\ fl l SUBSET fl m ==> l inseg m ORDINAL_FL_UNIQUE = |- !l m. ordinal l /\ ordinal m /\ fl l = fl m ==> l = m TUKEY = |- !s. ~(s = {}) /\ (!t. (!c. FINITE c /\ c SUBSET t ==> c IN s) <=> t IN s) ==> (?u. u IN s /\ (!v. v IN s /\ u SUBSET v ==> u = v)) WF_INSEG_WOSET = |- WF (\x y. woset x /\ woset y /\ x inseg y /\ ~(x = y)) WOSET_INSEG_ORDINAL = |- woset (\(x,y). ordinal x /\ ordinal y /\ x inseg y) WOSET_WF = |- !l. woset l <=> WF (\x y. l(x,y) /\ ~(x = y)) /\ (!x y. fl l x /\ fl l y ==> l(x,y) \/ l(y,x)) Fri 10th Jul 2015 wf.ml Added the simple fact that a wellfounded relation is antisymmetric WF_ANTISYM = |- !(<<) x y. WF (<<) ==> ~(x << y /\ y << x) Tue 30th Jun 2015 sets.ml Added a variety of set theorems, mostly basic properties of CROSS that are near-clones of existing theorems for PCROSS, many with the same proof. CROSS_DIFF = |- (!s t u. s CROSS (t DIFF u) = s CROSS t DIFF s CROSS u) /\ (!s t u. (s DIFF t) CROSS u = s CROSS u DIFF t CROSS u) CROSS_EQ = |- !s s' t t'. s CROSS t = s' CROSS t' <=> (s = {} \/ t = {}) /\ (s' = {} \/ t' = {}) \/ s = s' /\ t = t' CROSS_INTER = |- (!s t u. s CROSS (t INTER u) = s CROSS t INTER s CROSS u) /\ (!s t u. (s INTER t) CROSS u = s CROSS u INTER t CROSS u) CROSS_MONO = |- !s t s' t'. s SUBSET s' /\ t SUBSET t' ==> s CROSS t SUBSET s' CROSS t' CROSS_UNION = |- (!s t u. s CROSS (t UNION u) = s CROSS t UNION s CROSS u) /\ (!s t u. (s UNION t) CROSS u = s CROSS u UNION t CROSS u) CROSS_UNIONS = |- (!s f. s CROSS UNIONS f = UNIONS {s CROSS t | t IN f}) /\ (!f t. UNIONS f CROSS t = UNIONS {s CROSS t | s IN f}) CROSS_UNIONS_UNIONS = |- !f g. UNIONS f CROSS UNIONS g = UNIONS {s CROSS t | s IN f /\ t IN g} EXISTS_IN_CROSS = |- !P s t. (?z. z IN s CROSS t /\ P z) <=> (?x y. x IN s /\ y IN t /\ P (x,y)) FORALL_IN_CROSS = |- !P s t. (!z. z IN s CROSS t ==> P z) <=> (!x y. x IN s /\ y IN t ==> P (x,y)) IMAGE_FST_CROSS = |- !s t. IMAGE FST (s CROSS t) = (if t = {} then {} else s) IMAGE_SND_CROSS = |- !s t. IMAGE SND (s CROSS t) = (if s = {} then {} else t) INTER_CROSS = |- !s s' t t'. s CROSS t INTER s' CROSS t' = (s INTER s') CROSS (t INTER t') SUBSET_CROSS = |- !s t s' t'. s CROSS t SUBSET s' CROSS t' <=> s = {} \/ t = {} \/ s SUBSET s' /\ t SUBSET t' UNIONS_UNIV = |- UNIONS (:A->bool) = (:A) Fri 12th Jun 2015 sets.ml Added two useful rewrites for sup and inf on finite sets from Marco Maggesi. Note that they don't a priori require the set to be finite, so they are easy to apply. INF_INSERT_INSERT = |- !a b s. inf (b INSERT a INSERT s) = inf (min a b INSERT s) SUP_INSERT_INSERT = |- !a b s. sup (b INSERT a INSERT s) = sup (max a b INSERT s) Fri 29th May 2015 Library/card.ml Added two more theorems about representing countable sets as images: COUNTABLE_AS_IMAGE_NUM_SUBSET = |- !s. COUNTABLE s <=> (?f k. s = IMAGE f k) COUNTABLE_AS_INJECTIVE_IMAGE_SUBSET = |- !s. COUNTABLE s <=> (?f k. s = IMAGE f k /\ (!m n. m IN k /\ n IN k /\ f m = f n ==> m = n)) Fri 15th May 2015 pair.ml, arith.ml, sets.ml Added a few useful theorems from Marco Maggesi: CHOICE_PAIRED_THM = |- !P Q. (?x y. P x y) /\ (!x y. P x y ==> Q (x,y)) ==> Q (@(x,y). P x y) CHOICE_UNPAIR_THM = |- !P. (@(x,y). P x y) = (@p. P (FST p) (SND p)) INTERS_SUBSET = |- !u s. ~(u = {}) /\ (!t. t IN u ==> t SUBSET s) ==> INTERS u SUBSET s LAMBDA_UNPAIR_THM = |- !f. (\(x,y). f x y) = (\p. f (FST p) (SND p)) LE_INDUCT = |- !P. (!m. P m m) /\ (!m n. m <= n /\ P m n ==> P m (SUC n)) ==> (!m n. m <= n ==> P m n) Fri 24th Apr 2015 sets.ml Slightly sharpened PAIRWISE_IMP to |- !P Q s. pairwise P s /\ (!x y. x IN s /\ y IN s /\ P x y /\ ~(x = y) ==> Q x y) ==> pairwise Q s and added one more trivial but useful theorem about "pairwise": PAIRWISE_CHAIN_UNIONS = |- !R c. (!s. s IN c ==> pairwise R s) /\ (!s t. s IN c /\ t IN c ==> s SUBSET t \/ t SUBSET s) ==> pairwise R (UNIONS c) Fri 17th Apr 2015 calc_rat.ml, int.ml Added a few equivalential forms of transitivity; the first few in particular can be useful for epsilon-delta arguments in analysis. REAL_LE_TRANS_LE = |- !x y:real. x <= y <=> (!z. y <= z ==> x <= z) REAL_LE_TRANS_LTE = |- !x y:real. x <= y <=> (!z. y < z ==> x <= z) REAL_LE_TRANS_LT = |- !x y:real. x <= y <=> (!z. y < z ==> x < z) INT_LE_TRANS_LE = |- !x y:int x <= y <=> (!z. y <= z ==> x <= z) INT_LE_TRANS_LT = |- !x y:int. x <= y <=> (!z. y < z ==> x < z) Mon 30th Mar 2015 sets.ml Added a simple lemma about "pairwise": PAIRWISE_UNION = |- !R s t. pairwise R (s UNION t) <=> pairwise R s /\ pairwise R t /\ (!x y. x IN s DIFF t /\ y IN t DIFF s ==> R x y /\ R y x) Fri 20th Mar 2015 sets.ml, Library/card.ml Added a few convenient rewrites extending existing ones with an injectivity assumption: EXISTS_COUNTABLE_SUBSET_IMAGE_INJ = |- !P f s. (?t. COUNTABLE t /\ t SUBSET IMAGE f s /\ P t) <=> (?t. COUNTABLE t /\ t SUBSET s /\ (!x y. x IN t /\ y IN t ==> (f x = f y <=> x = y)) /\ P (IMAGE f t)) EXISTS_FINITE_SUBSET_IMAGE_INJ = |- !P f s. (?t. FINITE t /\ t SUBSET IMAGE f s /\ P t) <=> (?t. FINITE t /\ t SUBSET s /\ (!x y. x IN t /\ y IN t ==> (f x = f y <=> x = y)) /\ P (IMAGE f t)) EXISTS_SUBSET_IMAGE_INJ = |- !P f s. (?t. t SUBSET IMAGE f s /\ P t) <=> (?t. t SUBSET s /\ (!x y. x IN t /\ y IN t ==> (f x = f y <=> x = y)) /\ P (IMAGE f t)) FORALL_COUNTABLE_SUBSET_IMAGE_INJ = |- !P f s. (!t. COUNTABLE t /\ t SUBSET IMAGE f s ==> P t) <=> (!t. COUNTABLE t /\ t SUBSET s /\ (!x y. x IN t /\ y IN t ==> (f x = f y <=> x = y)) ==> P (IMAGE f t)) FORALL_FINITE_SUBSET_IMAGE_INJ = |- !P f s. (!t. FINITE t /\ t SUBSET IMAGE f s ==> P t) <=> (!t. FINITE t /\ t SUBSET s /\ (!x y. x IN t /\ y IN t ==> (f x = f y <=> x = y)) ==> P (IMAGE f t)) FORALL_SUBSET_IMAGE_INJ = |- !P f s. (!t. t SUBSET IMAGE f s ==> P t) <=> (!t. t SUBSET s /\ (!x y. x IN t /\ y IN t ==> (f x = f y <=> x = y)) ==> P (IMAGE f t)) SUBSET_IMAGE_INJ = |- !f s t. s SUBSET IMAGE f t <=> (?u. u SUBSET t /\ (!x y. x IN u /\ y IN u ==> (f x = f y <=> x = y)) /\ s = IMAGE f u) Wed 11th Mar 2015 sets.ml Added a couple of trivial but convenient sup/inf theorems: ELEMENT_LE_SUP = |- !s a. (?b. !x. x IN s ==> x <= b) /\ a IN s ==> a <= sup s INF_LE_ELEMENT = |- !s a. (?b. !x. x IN s ==> b <= x) /\ a IN s ==> inf s <= a Fri 6th Mar 2015 sets.ml Added PAIRWISE_AND = |- !R R' s. pairwise R s /\ pairwise R' s <=> pairwise (\x y. R x y /\ R' x y) s Fri 6th Feb 2015 iterate.ml Added two rather trivial but occasionally useful results: NSUM_MUL_BOUND = |- !a b s. FINITE s ==> nsum s (\i. a i * b i) <= nsum s a * nsum s b SUM_MUL_BOUND;; |- !a b s. FINITE s /\ (!i. i IN s ==> &0 <= a i /\ &0 <= b i) ==> sum s (\i. a i * b i) <= sum s a * sum s b Fri 30th Jan 2015 sets.ml Added the following, which is convenient to avoid manual reshuffling around FORALL_IN_IMAGE: FORALL_IN_IMAGE_2 = |- !f P s. (!x y. x IN IMAGE f s /\ y IN IMAGE f s ==> P x y) <=> (!x y. x IN s /\ y IN s ==> P (f x) (f y)) Fri 23rd Jan 2015 Library/card.ml Added CARD_MUL_FINITE_EQ = |- !s t. FINITE (s *_c t) <=> s = {} \/ t = {} \/ FINITE s /\ FINITE t Fri 23rd Jan 2015 sets.ml Added three more simple theorems: CROSS_EMPTY = |- (!s. s CROSS {} = {}) /\ (!t. {} CROSS t = {}) FINITE_CROSS_EQ = |- !s t. FINITE (s CROSS t) <=> s = {} \/ t = {} \/ FINITE s /\ FINITE t PAIRWISE_IMP = |- !P Q s. pairwise P s /\ (!x y. P x y /\ ~(x = y) ==> Q x y) ==> pairwise Q s Mon 12th Jan 2015 sets.ml Added another related theorem to the previous addition: INTER_UNIONS_PAIRWISE_DISJOINT = |- !s t. pairwise DISJOINT (s UNION t) ==> UNIONS s INTER UNIONS t = UNIONS (s INTER t) Fri 9th Jan 2015 sets.ml Added this fairly obvious theorem: DIFF_UNIONS_PAIRWISE_DISJOINT = |- !s t. pairwise DISJOINT s /\ t SUBSET s ==> UNIONS s DIFF UNIONS t = UNIONS(s DIFF t) Sat 27th Dec 2014 Quaternions/* Added a substantial new theory of quaternions from Marco Maggesi. This includes basic operations, automatic normalization, quaternionic analysis and the relation of quaternions to orthogonal transformations. Sat 13th Dec 2014 theorems.ml Added CLAIM_TAC from Marco Maggesi. This combines SUBGOAL_THEN and DESTRUCT_TAC in a natural way so one can make a labelled claim with a subproof. Sat 6th Dec 2014 IEEE/* [new directory] Added an extensive formal theory of IEEE floating-point arithmetic, written by Charlie Jacobsen. This includes theories of fixed-point numbers, generalized floating-point numbers (with any even radix >= 2 and any precision >= 2) and then the properties of the usual IEEE floating-point numbers. Sun 30th Nov 2014 lists.ml, Help/dest_char.doc [new file], Help/mk_char.doc [new file], Help/dest_string.doc [new file], Help/mk_string.doc [new file], Help/CHAR_EQ_CONV.doc [new file], Help/STRING_EQ_CONV.doc [new file] Added a systematic set of syntax constructors, destructors and comparison conversions for characters and for strings (== char lists), written by Marco Maggesi: dest_char mk_char dest_string mk_string CHAR_EQ_CONV STRING_EQ_CONV Sun 30th Nov 2014 parser.ml Fixed a bug in the parsing of HOL strings where the special handling of a quotation with just a single identifier (designed to allow one to omit the parantheses round infixes in this special case) was accidentally catching strings as well. Sun 23rd Nov 2014 sets.ml Added this, a natural dual to the existing INTERS_OVER_UNIONS UNIONS_OVER_INTERS = |- !f s. UNIONS {INTERS (f x) | x IN s} = INTERS {UNIONS {g x | x IN s} | g | !x. x IN s ==> g x IN f x} Fri 7th Nov 2014 Examples/misiurewicz.ml [new file] Added a new example, a formalization (suggested on the FOM list by Lasse Rempe-Gillen) of Misiurewicz's original proof that the complex exponential map is topologically transitive. See "On iterates of e^z", Ergodic Theory and Dynamical Systems, vol 1, pp. 103-106, 1981. Fri 7th Nov 2014 Library/iter.ml Added ITER_ADD_POINTLESS : thm = |- !m n. ITER (m + n) f = ITER m f o ITER n f Fri 7th Nov 2014 sets.ml Added FINITE_SUBSET_NUMSEG = |- !s. FINITE s <=> (?n. s SUBSET 0..n) Fri 31st Oct 2014 sets.ml Added the following, both stronger and more convenient than the existing INFINITE_IMAGE_INJ (which is kept for compatibility): INFINITE_IMAGE |- !f s. INFINITE s /\ (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> INFINITE (IMAGE f s) Fri 31st Oct 2014 theorems.ml, ind_defs.ml Added a general form of "without loss of generality" lemmas like REAL_WLOG_LE, as well as reshuffling other theorems from ind_defs.ml to theorems.ml. WLOG_RELATION = |- !R P. (!x y. P x y ==> P y x) /\ (!x y. R x y \/ R y x) /\ (!x y. R x y ==> P x y) ==> (!x y. P x y) Sun 12th Oct 2014 theorems.ml, Help/DESTRUCT_TAC/doc, Help/FIX_TAX.doc, Help/HYP_TAC.doc [new file] Added a new tactic HYP_TAC from Marco Maggesi. This fits into the existing suite of tatics (FIX_TAC, DESTRUCT_TAC etc.) for manipulations via brief strings, now applying to a named hypothesis. Also added a few improvements to the existing suite and their documentation. Sun 12th Oct 2014 Library/integer.ml Removed the theorem INT_CONG_0, which as pointed out by Marco Maggesi is the same as INT_CONG_0_DIVIDES except for the order of quantification. The longer name was kept because it fits the natural number case which is called CONG_DIVIDES_0, and it is actually used elsewhere. Thu 9th Oct 2014 iterate.ml Added variants of "grouping" theorems for sums, not requiring an explicit function: NSUM_GROUP_RELATION = |- !R g s t. FINITE s /\ (!x. x IN s ==> (?!y. y IN t /\ R x y)) ==> nsum t (\y. nsum {x | x IN s /\ R x y} g) = nsum s g SUM_GROUP_RELATION = |- !R g s t. FINITE s /\ (!x. x IN s ==> (?!y. y IN t /\ R x y)) ==> sum t (\y. sum {x | x IN s /\ R x y} g) = sum s g Sat 4th Oct 2014 Functionspaces/README, Functionspaces/make.ml, Functionspaces/utils.ml, Functionspaces/cfunspace.ml [new files], Jordan/float.ml Added a new library by Mohamed Yousri Mahmoud and Vincent Aravantinos of vector spaces of complex functions, building on top of the Multivariate theories. Sat 4th Oct 2014 system.ml Added another form to the quotation expander so that any quotation ending in a colon, i.e. of the form `....:`, will just be parsed as an (escaped) string. This is convenient for use with the new Library/q.ml file so that potential issues with capturing backslashes in plain strings can be avoided. It seems quite an intuitive syntax since one is then expecting the system to infer the type from context. Tweaked cases in Jordan/float.ml where there was already a colon at the end of a quotation, just adding spaces around the body. Fri 26th Sep 2014 sets.ml Added the following, a natural complement to IMAGE_UNIONS and IMAGE_INTER: IMAGE_INTERS = |- !f s. ~(s = {}) /\ (!x y. x IN UNIONS s /\ y IN UNIONS s /\ f x = f y ==> x = y) ==> IMAGE f (INTERS s) = INTERS(IMAGE (IMAGE f) s) Fri 19th Sep 2014 Library/q.ml [new file] Added a file from Vincent Aravantinos with Pa.xxx variants of many standard HOL Light functions, taking strings instead of terms and inferring types appropriately. This is similar to the Q module in HOL4. Thu 18th Sep 2014 sets.ml Added a couple more set theory rewrites: FORALL_SUBSET_INSERT = |- !a t. (!s. s SUBSET a INSERT t ==> P s) <=> (!s. s SUBSET t ==> P s /\ P (a INSERT s)) EXISTS_SUBSET_INSERT = |- !a t. (?s. s SUBSET a INSERT t /\ P s) <=> (?s. s SUBSET t /\ (P s \/ P (a INSERT s))) Tue 16th Sep 2014 Library/permutations.ml Added a couple more theorems about permutations: PERMUTES_ID = |- !s. (\x. x) permutes s REAL_SGN_SIGN = |- !p. real_sgn (sign p) = sign p Tue 9th Sep 2014 Library/card.ml Added a few more general lemma about cardinal operations CARD_EMPTY_LE = |- !s. {} <=_c s CARD_EQ_REFL_IMP = |- !s t. s = t ==> s =_c t CARD_LE_FINITE_INFINITE = |- !s t. FINITE s /\ INFINITE t ==> s <=_c t CARD_MUL_C = |- !s t. FINITE s /\ FINITE t ==> CARD (s *_c t) = CARD s * CARD t CARD_SING_LE = |- !a s. {a} <=_c s <=> ~(s = {}) LE_C_IMAGE = |- !s t. s <=_c t <=> s = {} \/ (?f. IMAGE f t = s) together with a definition of cardinal exponentiation and its basic properties: exp_c = |- !s t. s ^_c t = {f | (!x. x IN t ==> f x IN s) /\ (!x. ~(x IN t) ==> f x = (@y. F))} CARD_EQ_COUNTABLE_SUBSETS_SUBREAL = |- !s. INFINITE s /\ s <=_c (:real) ==> {t | t SUBSET s /\ COUNTABLE t} =_c (:real) CARD_EQ_FULLSIZE_POWERSET = |- !s. INFINITE s ==> {t | t SUBSET s /\ t =_c s} =_c {t | t SUBSET s} CARD_EQ_LIMITED_POWERSET = |- !s t. INFINITE s ==> (if t <=_c s then {k | k SUBSET s /\ k <=_c t} =_c s ^_c t else {k | k SUBSET s /\ k <=_c t} =_c (:bool) ^_c s) CARD_EQ_RESTRICTED_POWERSET = |- !s t. INFINITE s ==> {k | k SUBSET s /\ k =_c t} =_c (if t <=_c s then s ^_c t else {}) CARD_EXP_0 = |- !s c. s ^_c {} =_c {c}; CARD_EXP_ABSORB = |- !s t. INFINITE t /\ (:bool) <=_c s /\ s <=_c (:bool) ^_c t ==> s ^_c t =_c (:bool) ^_c t CARD_EXP_ADD = |- !s t u. s ^_c (t +_c u) =_c s ^_c t *_c s ^_c u CARD_EXP_C = |- !s t. FINITE s /\ FINITE t ==> CARD (s ^_c t) = CARD s EXP CARD t CARD_EXP_CANTOR = |- !s. s <_c (:bool) ^_c s CARD_EXP_CONG = |- !s s' t t'. s =_c s' /\ t =_c t' ==> s ^_c t =_c s' ^_c t' CARD_EXP_FINITE = |- !s t. FINITE s /\ FINITE t ==> FINITE (s ^_c t) CARD_EXP_GRAPH = |- !s t. s ^_c t =_c {R | (!x y. R x y ==> x IN t /\ y IN s) /\ (!x. x IN t ==> (?!y. R x y))} CARD_EXP_GRAPH_PAIRED = |- !s t. s ^_c t =_c {R | (!x y. R (x,y) ==> x IN t /\ y IN s) /\ (!x. x IN t ==> (?!y. R (x,y)))} CARD_EXP_MUL = |- !s t u. s ^_c (t *_c u) =_c s ^_c t ^_c u CARD_EXP_POWERSET = |- !s. (:bool) ^_c s =_c {t | t SUBSET s} CARD_EXP_SING = |- !s b. s ^_c {b} =_c s CARD_EXP_UNIV = |- (:A) ^_c (:B) = (:B->A) CARD_EXP_ZERO = |- !s c. {} ^_c s =_c (if s = {} then {c} else {}) CARD_LE_EXP = |- !s s' t t'. ~(s = {}) /\ s <=_c s' /\ t <=_c t' ==> s ^_c t <=_c s' ^_c t' CARD_LE_EXP_LEFT = |- !s s' t. s <=_c s' ==> s ^_c t <=_c s' ^_c t CARD_LE_EXP_RIGHT = |- !s t t'. ~(s = {}) /\ t <=_c t' ==> s ^_c t <=_c s ^_c t' CARD_MUL_EXP = |- !s t u. (s *_c t) ^_c u =_c s ^_c u *_c t ^_c u Tue 9th Sep 2014 Library/products.ml Added two straightforward "finite support" generalizations of existing theorems about products: NPRODUCT_MUL_GEN = |- !f g s. FINITE {x | x IN s /\ ~(f x = 1)} /\ FINITE {x | x IN s /\ ~(g x = 1)} ==> nproduct s (\x. f x * g x) = nproduct s f * nproduct s g PRODUCT_MUL_GEN = |- !f g s. FINITE {x | x IN s /\ ~(f x = &1)} /\ FINITE {x | x IN s /\ ~(g x = &1)} ==> product s (\x. f x * g x) = product s f * product s g Tue 19th Aug 2014 Library/floor.ml Added NONNEGATIVE_INTEGER = |- !x. integer x /\ &0 <= x <=> ?n. x = &n NONPOSITIVE_INTEGER = |- !x. integer x /\ x <= &0 <=> ?n. x = -- &n NONPOSITIVE_INTEGER_ALT = |- !x. integer x /\ x <= &0 <=> ?n. x + &n = &0 REAL_FLOOR_NEG = |- !x. floor(--x) = if integer x then --x else --(floor x + &1) FRAC_NEG = |- !x. frac(--x) = if integer x then &0 else &1 - frac x Tue 19th Aug 2014 iterate.ml Added REAL_OF_NUM_SUM_GEN = |- !f s. FINITE {i | i IN s /\ ~(f i = 0)} ==> &(nsum s f) = sum s (\x. &(f x))`, Tue 19th Aug 2014 real.ml Added REAL_OF_NUM_SUB_CASES = |- `!m n. &m - &n = if n <= m then &(m - n) else -- &(n - m) Sun 13th Jul 2014 Formal_ineqs/* [new files] Added Alexey Solovyev's formal inequality prover (used extensively in the Flyspeck project) to the HOL Light distribution. Thu 10th Jul 2014 Library/products.ml Added a couple more trivial theorems about products: NPRODUCT_DELTA = |- !s a. nproduct s (\x. if x = a then b else 1) = (if a IN s then b else 1) PRODUCT_DELTA = |- !s a. product s (\x. if x = a then b else &1) = (if a IN s then b else &1) Tue 1st Jul 2014 Library/floor.ml Added clauses for "max" and "min" to INTEGER_CLOSED, so it is now: INTEGER_CLOSED = |- (!n. integer (&n)) /\ (!x y. integer x /\ integer y ==> integer (x + y)) /\ (!x y. integer x /\ integer y ==> integer (x - y)) /\ (!x y. integer x /\ integer y ==> integer (x * y)) /\ (!x r. integer x ==> integer (x pow r)) /\ (!x. integer x ==> integer (--x)) /\ (!x. integer x ==> integer (abs x)) /\ (!x y. integer x /\ integer y ==> integer (max x y)) /\ (!x y. integer x /\ integer y ==> integer (min x y)) Tue 1st Jul 2014 sets.ml Added CROSS_UNIV = |- (:A) CROSS (:B) = (:A#B) Tue 27th May 2014 Library/permutations.ml Added three more lemmas about permutations: CARD_EVEN_PERMUTATIONS = |- !s. FINITE s /\ 2 <= CARD s ==> 2 * CARD {p | p permutes s /\ evenperm p} = FACT (CARD s) PERMUTES_INVOLUTION = |- !p s. (!x. p (p x) = x) /\ (!x. ~(x IN s) ==> p x = x) ==> p permutes s SIGN_INVOLUTION = |- !p s. FINITE s /\ (!x. p (p x) = x) /\ (!x. ~(x IN s) ==> p x = x) ==> sign p = -- &1 pow (CARD {x | ~(p x = x)} DIV 2) Tue 27th May 2014 Library/floor.ml Added the following theorem, a natural counterpart for FLOOR_DIV_DIV FRAC_DIV_MOD = |- !m n. ~(n = 0) ==> frac(&m / &n) = &(m MOD n) / &n Mon 12th May 2014 Examples/gcdrecurrence.ml [new file] Added a new file with the cute result that the gcd operation "commutes" with certain integer sequences, with applications to the Mersenne numbers, the Fibonacci series, and solutions to the Pell equation. Sat 10th May 2014 int.ml, real.ml, Library/prime.ml, Library/integer.ml Added a few miscellaneous trivial theorems: DIVIDES_GCD_LEFT = |- !m n. m divides n <=> gcd (m,n) = m DIVIDES_GCD_RIGHT = |- !m n. n divides m <=> gcd (m,n) = n INT_DIVIDES_ANTISYM_DIVISORS = |- !a b. a divides b /\ b divides a <=> (!d. d divides a <=> d divides b) INT_DIVIDES_GCD_LEFT = |- !m n. m divides n <=> gcd (m,n) = abs m INT_DIVIDES_GCD_RIGHT = |- !m n. n divides m <=> gcd (m,n) = abs n INT_EQ_SGN_ABS = |- !x y. x = y <=> int_sgn x = int_sgn y /\ abs x = abs y REAL_EQ_SGN_ABS = |- !x y. x = y <=> real_sgn x = real_sgn y /\ abs x = abs y Sat 10th May 2014 Examples/dickson.ml Added a proof of Dickson's Lemma, plus the Nash-Williams minimal bad sequence property in a bit more generality (for a "safety property" of sequences in the Lamport/Alpern/Schneider sense). Fri 9th May 2014 Library/card.ml Added COUNTABLE_SUBSET_NUM = |- !s:num->bool. COUNTABLE s Sun 27th Apr 2014 Library/permutations.ml Added PERMUTES_BIJECTIONS = |- !p q. (!x. x IN s ==> p x IN s) /\ (!x. ~(x IN s) ==> p x = x) /\ (!x. x IN s ==> q x IN s) /\ (!x. ~(x IN s) ==> q x = x) /\ (!x. p(q x) = x) /\ (!x. q(p x) = x) ==> p permutes s Sat 19th Apr 2014 sets.ml Added another clause to FORALL_IN_GSPEC and EXISTS_IN_GSPEC for the case of 4 bound variables, so they have now become: FORALL_IN_GSPEC = |- (!P f. (!z. z IN {f x | P x} ==> Q z) <=> (!x. P x ==> Q(f x))) /\ (!P f. (!z. z IN {f x y | P x y} ==> Q z) <=> (!x y. P x y ==> Q(f x y))) /\ (!P f. (!z. z IN {f w x y | P w x y} ==> Q z) <=> (!w x y. P w x y ==> Q(f w x y))) /\ (!P f. (!z. z IN {f v w x y | P v w x y} ==> Q z) <=> (!v w x y. P v w x y ==> Q(f v w x y))) EXISTS_IN_GSPEC = |- (!P f. (?z. z IN {f x | P x} /\ Q z) <=> (?x. P x /\ Q(f x))) /\ (!P f. (?z. z IN {f x y | P x y} /\ Q z) <=> (?x y. P x y /\ Q(f x y))) /\ (!P f. (?z. z IN {f w x y | P w x y} /\ Q z) <=> (?w x y. P w x y /\ Q(f w x y))) /\ (!P f. (?z. z IN {f v w x y | P v w x y} /\ Q z) <=> (?v w x y. P v w x y /\ Q(f v w x y))) Sat 19th Apr 2014 pair.ml Added "unpair" theorems: FORALL_UNPAIR_THM = |- (!x y. P x y) <=> (!z. P (FST z) (SND z)) EXISTS_UNPAIR_THM = |- (?x y. P x y) <=> (?z. P (FST z) (SND z)) Tue 15th Apr 2014 bool.ml Did a little more rationalization of the Boolean rule, recoding MP so that it produces precisely the union of the two assumption lists, as the documentation claims. Mon 14th Apr 2014 Examples/lucas_lehmer.ml, 100/sqrt.ml [new file] Added a new file containing a formalization of the Lucas-Lehmer test for primality of Mersenne numbers, and a derived rule for applying it in particular cases. This is quite a bit faster than the usual PRIME_CONV (in this special case 2^p-1), and also scales better because it doesn't rely on factorization at all. This uses as a lemma the fact that sqrt(3) is not rational, so also added a file with a trivial proof that sqrt(p) is irrational for prime p. Mon 14th Apr 2014 int.ml, Library/prime.ml, Library/iter.ml, Library/pocklington.ml, Library/primitive.ml Moved this very basic result out of the Library/prime.ml file into the main core DIVIDES_LE = |- !m n. m divides n ==> m <= n \/ n = 0 Added some theorems to Library/iter.ml about the "order" or "characteristic" in a general setting val ORDER_EXISTENCE_GEN : thm = |- !P f. P (f 0) /\ (!m n. P (f m) /\ ~(m = 0) ==> (P (f (m + n)) <=> P (f n))) ==> ?d. !n. P (f n) <=> d divides n val ORDER_EXISTENCE_ITER : thm = |- !R f z. R z z /\ (!x y. R x y ==> R y x) /\ (!x y z. R x y /\ R y z ==> R x z) /\ (!x y. R x y ==> R (f x) (f y)) ==> ?d. !n. R (ITER n f z) z <=> d divides n val ORDER_EXISTENCE_CARD : thm = |- !R f z k. FINITE {R (ITER n f z) | n IN (:num)} /\ CARD {R (ITER n f z) | n IN (:num)} <= k /\ R z z /\ (!x y. R x y ==> R y x) /\ (!x y z. R x y /\ R y z ==> R x z) /\ (!x y. R (f x) (f y) <=> R x y) ==> ?d. 0 < d /\ d <= k /\ (!n. R (ITER n f z) z <=> d divides n) val ORDER_EXISTENCE_FINITE : thm = |- !R f z. FINITE {R (ITER n f z) | n IN (:num)} /\ R z z /\ (!x y. R x y ==> R y x) /\ (!x y z. R x y /\ R y z ==> R x z) /\ (!x y. R (f x) (f y) <=> R x y) ==> ?d. 0 < d /\ (!n. R (ITER n f z) z <=> d divides n) Changed the definition of "order" in Library/pocklington.ml to use that as the basis, as well as adding three basic new theorems: EXP_ITER = |- !x n. x EXP n = ITER n (\y. x * y) 1 PHI_EQ_0 = |- !n. phi n = 0 <=> n = 0 ORDER_1 = |- !n. order n 1 = 1 Fri 11th Apr 2014 bool.ml Bill Richter pointed out that IMP_ANTISYM_RULE could sometimes give fewer assumptions than its documentation (with a simple union of the two lists) suggested. Fixed this, as well as a similar problem with CONJ, and a few related documentation issues also uncovered by Bill. Wed 9th Apr 2014 equal.ml Made a few minor optimizations to some functions here, using more pattern-matching in place of destructors and adding partial evaluation to AP_TERM. Fri 4th Apr 2014 arith.ml Added MOD_REFL = |- !n. ~(n = 0) ==> n MOD n = 0 Thu 3rd Apr 2014 nums.ml, calc_num.ml Completely rewrote the conversions NUM_SUC_CONV, NUM_ADD_CONV and NUM_MULT_CONV to improve efficiency, as well as the auxiliary syntax function mk_numeral. The new versions do addition in larger blocks for a decent constant factor improvement and the multiplication routine uses squaring as the core with quite tight coding and the use of 2-part and 3-part interpolation. Also rewrote the main relational conversions NUM_LT_CONV, NUM_LE_CONV and NUM_EQ_CONV to call addition directly, which is (now) much more efficient than what was used before, though it could still be optimized a bit more. Tue 1st Apr 2014 fusion.ml Made another rather minor optimization, this time making the check for type compatibility in the "vsubst" outer wrapper marginally more efficient. Thu 27th Mar 2014 fusion.ml Made a small change to the "type_of" function to use an explicit match in the Comb case instead of applying "hd" and "tl". Although perhaps marginally uglier, it is somewhat more efficient and gets two more library functions (hd and tl) out of the core. Sun 23rd Mar 2014 Examples/harmonicsum.ml [new file] Added a new file with a rather trivial but nice result that a non-trivial contiguous segment of the harmonic series never sums to an integer. Tue 18th Mar 2014 iterate.ml Added a definition of a "polynomial function" (R->R) with various basic closure properties. This is a sufficiently useful concept that it seems worth having in the core, not just in the analytical theories. polynomial_function = |- !p. polynomial_function p <=> (?m c. !x. p x = sum (0..m) (\i. c i * x pow i)) POLYNOMIAL_FUNCTION_ADD = |- !p q. polynomial_function p /\ polynomial_function q ==> polynomial_function (\x. p x + q x) POLYNOMIAL_FUNCTION_CONST = |- !c. polynomial_function (\x. c) POLYNOMIAL_FUNCTION_FINITE_ROOTS = |- !p a. polynomial_function p ==> (FINITE {x | p x = a} <=> ~(!x. p x = a)) POLYNOMIAL_FUNCTION_I = |- polynomial_function I POLYNOMIAL_FUNCTION_ID = |- polynomial_function (\x. x) POLYNOMIAL_FUNCTION_INDUCT = |- !P. P (\x. x) /\ (!c. P (\x. c)) /\ (!p q. P p /\ P q ==> P (\x. p x + q x)) /\ (!p q. P p /\ P q ==> P (\x. p x * q x)) ==> (!p. polynomial_function p ==> P p) POLYNOMIAL_FUNCTION_LMUL = |- !p c. polynomial_function p ==> polynomial_function (\x. c * p x) POLYNOMIAL_FUNCTION_MUL = |- !p q. polynomial_function p /\ polynomial_function q ==> polynomial_function (\x. p x * q x) POLYNOMIAL_FUNCTION_NEG = |- !p. polynomial_function (\x. --p x) <=> polynomial_function p POLYNOMIAL_FUNCTION_POW = |- !p n. polynomial_function p ==> polynomial_function (\x. p x pow n) POLYNOMIAL_FUNCTION_RMUL = |- !p c. polynomial_function p ==> polynomial_function (\x. p x * c) POLYNOMIAL_FUNCTION_SUB = |- !p q. polynomial_function p /\ polynomial_function q ==> polynomial_function (\x. p x - q x) POLYNOMIAL_FUNCTION_SUM = |- !s p. FINITE s /\ (!i. i IN s ==> polynomial_function (\x. p x i)) ==> polynomial_function (\x. sum s (p x)) POLYNOMIAL_FUNCTION_o = |- !p q. polynomial_function p /\ polynomial_function q ==> polynomial_function (p o q) Tue 18th Mar 2014 iterate.ml Added a conversion EXPAND_SUM_CONV to expand "sum (m..n) f" where m and n are particular numerals. Tue 18th Mar 2014 Library/binomial.ml Added one more lemma about binomial coefficients, a characteristic property of Appell sequences: APPELL_SEQUENCE = |- !c n x y. sum (0..n) (\k. &(binom (n,k)) * sum (0..k) (\l. &(binom (k,l)) * c l * x pow (k - l)) * y pow (n - k)) = sum (0..n) (\k. &(binom (n,k)) * c k * (x + y) pow (n - k)) Sun 16th Mar 2014 Multivariate/gamma.ml [new file] Added a new file giving a definition of the complex and real gamma functions and proofs of some basic properties. Mon 10th Mar 2014 Library/binomial.ml Added a couple more lemmas: BINOM_SYM = |- !n k. binom(n,n - k) = (if k <= n then binom(n,k) else 1) BINOM_MUL_SHIFT = |- !m n k. k <= m ==> binom(n,m) * binom(m,k) = binom(n,k) * binom(n - k,m - k) as well as putting a simple symmetry optimization into NUM_BINOM_CONV for the case of NUM_BINOM_CONV `binom(n,k)` where k is small. Sat 8th Mar 2014 printer.ml Added two improvements from Joe Hurd to the treatment of printing. First, user printers are generalized over the formatter so that they will respect the overall output formatter settings. Moreover, the "print_to_string" function now takes the margin into account. Fri 28th Feb 2014 Library/prime.ml Added a few more number-theory results, mainly concerning the "index": INDEX_REFL = |- !n. index n n = (if n <= 1 then 0 else 1) INDEX_EQ_0 = |- !p n. index p n = 0 <=> n = 0 \/ p = 1 \/ ~(p divides n) INDEX_TRIVIAL_BOUND = |- !n p. index p n <= n INDEX_DECOMPOSITION = |- !n p. ?m. p EXP index p n * m = n /\ (n = 0 \/ p = 1 \/ ~(p divides m)) INDEX_DECOMPOSITION_PRIME = |- !n p. prime p ==> (?m. p EXP index p n * m = n /\ (n = 0 \/ coprime (p,m))) DIVIDES_NSUM = |- !n f s. FINITE s /\ (!i. i IN s ==> n divides f i) ==> n divides nsum s f Fri 28th Feb 2014 real.ml Added REAL_LT_POW_2 = |- |- !x. &0 < x pow 2 <=> ~(x = &0) Fri 28th Feb 2014 Library/products.ml Added a few more basic theorems about products: NPRODUCT_FACT = |- !n. nproduct (1..n) (\m. m) = FACT n NPRODUCT_PAIR = |- !f m n. nproduct (2 * m..2 * n + 1) f = nproduct (m..n) (\i. f (2 * i) * f (2 * i + 1)) NPRODUCT_DELETE = |- !f s a. FINITE s /\ a IN s ==> f a * nproduct (s DELETE a) f = nproduct s f PRODUCT_PAIR = |- !f m n. product (2 * m..2 * n + 1) f = product (m..n) (\i. f (2 * i) * f (2 * i + 1)) PRODUCT_DELETE = |- !f s a. FINITE s /\ a IN s ==> f a * product (s DELETE a) f = product s f Sat 22nd Feb 2014 impconv.ml, Help/CASE_REWRITE_TAC.doc, Help/HINT_EXISTS_TAC.doc, Help/IMP_REWRITE_TAC.doc, Help/SEQ_IMP_REWRITE_TAC.doc, Help/TARGET_REWRITE_TAC.doc [new files], hol.ml Added a new file from Vincent Aravantinos containing several powerful new tactics: IMP_REWRITE_TAC, TARGET_REWRITE_TAC, CASE_REWRITE_TAC, SEQ_IMP_REWRITE_TAC and HINT_EXISTS_TAC. Sat 22nd Feb 2014 arith.ml, Library/binary.ml, Library/prime.ml Added the rather trivial "finite Cantor's theorem", which is otherwise defined in several different places and is useful enough: LT_POW2_REFL = |- !n. n < 2 EXP n Deleted it from other places now that it's in the core. Sun 16th Feb 2014 real.ml Added the following additional properties of the signum function: INT_SGN_POW = |- !x n. int_sgn(x pow n) = int_sgn(x) pow n INT_SGN_POW_2 = |- !x. int_sgn(x pow 2) = int_sgn(abs x) INT_SGN_REAL_SGN = |- int_sgn(int_sgn x) = int_sgn x REAL_INV_SGN = |- !x. inv (real_sgn x) = real_sgn x REAL_SGN_POW = |- !x n. real_sgn(x pow n) = real_sgn(x) pow n REAL_SGN_POW_2 = |- !x. real_sgn(x pow 2) = real_sgn(abs x) REAL_SGN_REAL_SGN = |- real_sgn(real_sgn x) = real_sgn x Sun 16th Feb 2014 Library/prime.ml Added a conversion LCM_CONV to compute LCMs of particular numerals, plus a few more theorems including a definition of "index": PRIMEPOW_DIVIDES_PROD = |- !p k m n. prime p /\ (p EXP k) divides (m * n) ==> ?i j. (p EXP i) divides m /\ (p EXP j) divides n /\ k = i + j FINITE_EXP_LE = |- !P p n. 2 <= p ==> FINITE {j | P j /\ p EXP j <= n} FINITE_INDICES = |- !P p n. 2 <= p /\ ~(n = 0) ==> FINITE {j | P j /\ p EXP j divides n} index_def = |- index p n = if p <= 1 \/ n = 0 then 0 else CARD {j | 1 <= j /\ p EXP j divides n} INDEX_0 = |- !p. index p 0 = 0 PRIMEPOW_DIVIDES_INDEX = |- !n p k. p EXP k divides n <=> n = 0 \/ p = 1 \/ k <= index p n LE_INDEX = |- !n p k. k <= index p n <=> (n = 0 \/ p = 1 ==> k = 0) /\ p EXP k divides n INDEX_1 = |- !p. index p 1 = 0 INDEX_MUL = |- !m n. prime p /\ ~(m = 0) /\ ~(n = 0) ==> index p (m * n) = index p m + index p n INDEX_EXP = |- !p n k. prime p ==> index p (n EXP k) = k * index p n INDEX_FACT = |- !p n. prime p ==> index p (FACT n) = nsum(1..n) (\m. index p m) INDEX_FACT_ALT = |- !p n. prime p ==> index p (FACT n) = nsum {j | 1 <= j /\ p EXP j <= n} (\j. n DIV (p EXP j)) INDEX_FACT_UNBOUNDED = |- !p n. prime p ==> index p (FACT n) = nsum {j | 1 <= j} (\j. n DIV (p EXP j)) PRIMEPOW_DIVIDES_FACT = |- !p n k. prime p ==> (p EXP k divides FACT n <=> k <= nsum {j | 1 <= j /\ p EXP j <= n} (\j. n DIV (p EXP j))) Sun 16th Feb 2014 iterate.ml Added a slightly more delicate "finite support" vesion of sum comparison: NSUM_LE_GEN = |- !f g s. (!x. x IN s ==> f x <= g x) /\ FINITE {x | x IN s /\ ~(g x = 0)} ==> nsum s f <= nsum s g Wed 12th Feb 2014 Library/floor.ml Added INTEGER_DIV = |- !m n. integer(&m / &n) <=> n = 0 \/ n divides m Wed 12th Feb 2014 int.ml, Library/prime.ml Moved the theorem "divides" out of Library/prime.ml into the core, since it seems silly not to have such a basic fact available. Wed 12th Feb 2014 Library/binomial.ml Added BINOM_0 = |- !n. binom(0,n) = if n = 0 then 1 else 0 BINOM_GE_TOP = |- !m n. 1 <= m /\ m < n ==> n <= binom(n,m) Wed 12th Feb 2014 Library/products.ml Added NPRODUCT_SUPPORT = |- !f s. nproduct (support ( * ) f s) f = nproduct s f NPRODUCT_SUPERSET = |- !f u v. u SUBSET v /\ (!x. x IN v /\ ~(x IN u) ==> f(x) = 1) ==> nproduct v f = nproduct u f` PRODUCT_SUPPORT = |- !f s. product (support ( * ) f s) f = product s f PRODUCT_SUPERSET = |- !f:A->real u v. u SUBSET v /\ (!x. x IN v /\ ~(x IN u) ==> f(x) = &1) ==> product v f = product u f Wed 12th Feb 2014 iterate.ml Added a couple of convenient theorems about when sums are stricly positive: NSUM_POS_LT_ALL = |- !s f. FINITE s /\ ~(s = {}) /\ (!i. i IN s ==> 0 < f i) ==> 0 < nsum s f SUM_POS_LT_ALL = |- !s f. FINITE s /\ ~(s = {}) /\ (!i. i IN s ==> &0 < f i) ==> &0 < sum s f Thu 6th Feb 2014 int.ml Added INT_LE_DISCRETE = |- !x y:int. x <= y <=> x < y + &1 Thu 6th Feb 2014 Library/prime.ml Added a definition of lcm to this file, and a suite of theorems mainly about that. Also added "symmetric" clauses to GCD_0 and GCD_1. DIVIDES_LCM = |- !m n r. r divides m \/ r divides n ==> r divides lcm (m,n) DIVIDES_LCM_GCD = |- !m n d. d divides lcm (m,n) <=> d * gcd (m,n) divides m * n DIVISORS_EQ = |- !m n. m = n <=> (!d. d divides m <=> d divides n) GCD_LCM_DISTRIB = |- !a b c. gcd (a,lcm (b,c)) = lcm (gcd (a,b),gcd (a,c)) lcm = |- !m n. lcm (m,n) = (if m * n = 0 then 0 else (m * n) DIV gcd (m,n)) LCM = |- !m n. m divides lcm (m,n) /\ n divides lcm (m,n) /\ (!d. m divides d /\ n divides d ==> lcm (m,n) divides d) LCM_0 = |- (!n. lcm (0,n) = 0) /\ (!n. lcm (n,0) = 0) LCM_1 = |- (!n. lcm (1,n) = n) /\ (!n. lcm (n,1) = n) LCM_ASSOC = |- !m n p. lcm (m,lcm (n,p)) = lcm (lcm (m,n),p) LCM_DIVIDES = |- !m n d. lcm (m,n) divides d <=> m divides d /\ n divides d LCM_EQ = |- !x y u v. (!d. x divides d /\ y divides d <=> u divides d /\ v divides d) ==> lcm (x,y) = lcm (u,v) LCM_EXP = |- !n a b. lcm (a EXP n,b EXP n) = lcm (a,b) EXP n LCM_GCD_DISTRIB = |- !a b c. lcm (a,gcd (b,c)) = gcd (lcm (a,b),lcm (a,c)) LCM_LMUL = |- !a b c. lcm (c * a,c * b) = c * lcm (a,b) LCM_MULTIPLE = |- !a b. lcm (b,a * b) = a * b LCM_REFL = |- !n. lcm (n,n) = n LCM_RMUL = |- !a b c. lcm (a * c,b * c) = c * lcm (a,b) LCM_SYM = |- !m n. lcm (m,n) = lcm (n,m) LCM_UNIQUE = |- !d m n. m divides d /\ n divides d /\ (!e. m divides e /\ n divides e ==> d divides e) <=> d = lcm (m,n) LCM_ZERO = |- !m n. lcm (m,n) = 0 <=> m = 0 \/ n = 0 MULTIPLES_EQ = |- !m n. m = n <=> (!d. m divides d <=> n divides d) PRIMEPOW_DIVIDES_LCM = |- !m n p k. prime p ==> (p EXP k divides lcm (m,n) <=> p EXP k divides m \/ p EXP k divides n) PRIMEPOW_DIVISORS_DIVIDES = |- !m n. m divides n <=> (!p k. prime p /\ p EXP k divides m ==> p EXP k divides n) PRIMEPOW_DIVISORS_EQ = |- !m n. m = n <=> (!p k. prime p ==> (p EXP k divides m <=> p EXP k divides n)) Wed 5th Feb 2014 Library/products.ml Added natural number products to "Library/products.ml", with all the analogous theorems that are meaningful for their real counterparts: nproduct = |- nproduct = iterate (*) NPRODUCT_ADD_SPLIT = |- !f m n p. m <= n + 1 ==> nproduct (m..n + p) f = nproduct (m..n) f * nproduct (n + 1..n + p) f NPRODUCT_CLAUSES = |- (!f. nproduct {} f = 1) /\ (!x f s. FINITE s ==> nproduct (x INSERT s) f = (if x IN s then nproduct s f else f x * nproduct s f)) NPRODUCT_CLAUSES_LEFT = |- !f m n. m <= n ==> nproduct (m..n) f = f m * nproduct (m + 1..n) f NPRODUCT_CLAUSES_NUMSEG = |- (!m. nproduct (m..0) f = (if m = 0 then f 0 else 1)) /\ (!m n. nproduct (m..SUC n) f = (if m <= SUC n then nproduct (m..n) f * f (SUC n) else nproduct (m..n) f)) NPRODUCT_CLAUSES_RIGHT = |- !f m n. 0 < n /\ m <= n ==> nproduct (m..n) f = nproduct (m..n - 1) f * f n NPRODUCT_CLOSED = |- !P f s. P 1 /\ (!x y. P x /\ P y ==> P (x * y)) /\ (!a. a IN s ==> P (f a)) ==> P (nproduct s f) NPRODUCT_CONST = |- !c s. FINITE s ==> nproduct s (\x. c) = c EXP CARD s NPRODUCT_CONST_NUMSEG = |- !c m n. nproduct (m..n) (\x. c) = c EXP ((n + 1) - m) NPRODUCT_CONST_NUMSEG_1 = |- !c n. nproduct (1..n) (\x. c) = c EXP n NPRODUCT_EQ = |- !f g s. (!x. x IN s ==> f x = g x) ==> nproduct s f = nproduct s g NPRODUCT_EQ_0 = |- !f s. FINITE s ==> (nproduct s f = 0 <=> (?x. x IN s /\ f x = 0)) NPRODUCT_EQ_0_NUMSEG = |- !f m n. nproduct (m..n) f = 0 <=> (?x. m <= x /\ x <= n /\ f x = 0) NPRODUCT_EQ_1 = |- !f s. (!x. x IN s ==> f x = 1) ==> nproduct s f = 1 NPRODUCT_EQ_1_NUMSEG = |- !f m n. (!i. m <= i /\ i <= n ==> f i = 1) ==> nproduct (m..n) f = 1 NPRODUCT_EQ_NUMSEG = |- !f g m n. (!i. m <= i /\ i <= n ==> f i = g i) ==> nproduct (m..n) f = nproduct (m..n) g NPRODUCT_IMAGE = |- !f g s. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> nproduct (IMAGE f s) g = nproduct s (g o f) NPRODUCT_LE = |- !f s. FINITE s /\ (!x. x IN s ==> 0 <= f x /\ f x <= g x) ==> nproduct s f <= nproduct s g NPRODUCT_LE_NUMSEG = |- !f m n. (!i. m <= i /\ i <= n ==> 0 <= f i /\ f i <= g i) ==> nproduct (m..n) f <= nproduct (m..n) g NPRODUCT_MUL = |- !f g s. FINITE s ==> nproduct s (\x. f x * g x) = nproduct s f * nproduct s g NPRODUCT_MUL_NUMSEG = |- !f g m n. nproduct (m..n) (\x. f x * g x) = nproduct (m..n) f * nproduct (m..n) g NPRODUCT_OFFSET = |- !f m p. nproduct (m + p..n + p) f = nproduct (m..n) (\i. f (i + p)) NPRODUCT_ONE = |- !s. nproduct s (\n. 1) = 1 NPRODUCT_POS_LT = |- !f s. FINITE s /\ (!x. x IN s ==> 0 < f x) ==> 0 < nproduct s f NPRODUCT_POS_LT_NUMSEG = |- !f m n. (!x. m <= x /\ x <= n ==> 0 < f x) ==> 0 < nproduct (m..n) f NPRODUCT_SING = |- !f x. nproduct {x} f = f x NPRODUCT_SING_NUMSEG = |- !f n. nproduct (n..n) f = f n NPRODUCT_UNION = |- !f s t. FINITE s /\ FINITE t /\ DISJOINT s t ==> nproduct (s UNION t) f = nproduct s f * nproduct t f REAL_OF_NUM_NPRODUCT = |- !f s. FINITE s ==> &(nproduct s f) = product s (\x. &(f x)) Thu 23rd Jan 2014 iterate.ml, Library/isum.ml Made the incompatible "improvement" of removing the finiteness hypothesis from SUM_POS_LE, as well as adding two theorems about how sums of natural numbers and integers collapse in the degenerate case of infinite support: NSUM_DEGENERATE = |- !f s. ~FINITE {x | x IN s /\ ~(f x = 0)} ==> nsum s f = 0 SUM_DEGENERATE = |- !f s. ~FINITE {x | x IN s /\ ~(f x = &0)} ==> sum s f = &0 Although I didn't explicitly change the file Library/isum.ml, the theorem ISUM_POS_LE automatically inherits the removal of the finiteness hypothesis. Thu 16th Jan 2014 sets.ml Added a few more basic properties of sup and inf: REAL_LE_SUP = |- !s a b y. y IN s /\ a <= y /\ (!x. x IN s ==> x <= b) ==> a <= sup s REAL_INF_LE = |- !s a b y. y IN s /\ y <= b /\ (!x. x IN s ==> a <= x) ==> inf s <= b REAL_SUP_LE_EQ = |- !s y. ~(s = {}) /\ (?b. !x. x IN s ==> x <= b) ==> (sup s <= y <=> (!x. x IN s ==> x <= y)) REAL_LE_INF_EQ = |- !s t. ~(s = {}) /\ (?b. !x. x IN s ==> b <= x) ==> (y <= inf s <=> (!x. x IN s ==> y <= x)) SUP_UNIQUE = |- !s b. (!c. (!x. x IN s ==> x <= c) <=> b <= c) ==> sup s = b INF_UNIQUE = |- !s b. (!c. (!x. x IN s ==> c <= x) <=> c <= b) ==> inf s = b SUP_UNION = |- !s t. ~(s = {}) /\ ~(t = {}) /\ (?b. !x. x IN s ==> x <= b) /\ (?c. !x. x IN t ==> x <= c) ==> sup (s UNION t) = max (sup s) (sup t) INF_UNION = |- !s t. ~(s = {}) /\ ~(t = {}) /\ (?b. !x. x IN s ==> b <= x) /\ (?c. !x. x IN t ==> c <= x) ==> inf (s UNION t) = min (inf s) (inf t) Thu 9th Jan 2014 class.ml, sets.ml Added a few theorems, two trivial but useful decomposition theorems for quantifiers over set unions, and two forms of the Axiom of Dependent Choice. These are a bit gratuitously overparameterized by "n", but this gives wider applicability and still works in the standard unparameterized case. FORALL_IN_UNION = |- !P s t. (!x. x IN s UNION t ==> P x) <=> (!x. x IN s ==> P x) /\ (!x. x IN t ==> P x) EXISTS_IN_UNION = |- !P s t. (?x. x IN s UNION t /\ P x) <=> (?x. x IN s /\ P x) \/ (?x. x IN t /\ P x) DEPENDENT_CHOICE_FIXED = |- !P R a. P 0 a /\ (!n x. P n x ==> (?y. P (SUC n) y /\ R n x y)) ==> ?f. f 0 = a /\ (!n. P n (f n)) /\ (!n. R n (f n) (f (SUC n))) DEPENDENT_CHOICE = |- !P R. (?a. P 0 a) /\ (!n x. P n x ==> (?y. P (SUC n) y /\ R n x y)) ==> ?f. (!n. P n (f n)) /\ (!n. R n (f n) (f (SUC n))) Mon 6th Jan 2014 holtest, holtest.mk [new file], holtest_parallel [new file] Added a new parallel version of holtest from Hendrik Tews, which is able to distribute tests over available processors/threads and is in general quite a bit faster. Also unified this and the old serial "holtest" over what to use as the HOL image: just use "hol-light" if it exists; make a "hol" executable if ckpt is available, otherwise just use ocaml directly. Fri 3rd Jan 2014 sets.ml, Permutation/morelist.ml Added a couple of simple lemmas about "list_of_set", for the simple case where there is no question about ordering: LIST_OF_SET_EMPTY = |- list_of_set {} = [] LIST_OF_SET_SING = |- !x. list_of_set {a} = [a] and a couple more related theorems in the Permutations library: LIST_UNIQ_APPEND = |- !l m. LIST_UNIQ (APPEND l m) <=> LIST_UNIQ l /\ LIST_UNIQ m /\ (!x. ~(MEM x l /\ MEM x m)) LIST_UNIQ_LIST_OF_SET = |- !s. FINITE s ==> LIST_UNIQ (list_of_set s) Sun 27th Oct 2013 class.ml Added a slightly more refined variant of SKOLEM_THM: SKOLEM_THM_GEN = |- !P s. (!x. P x ==> ?y. R x y) <=> (?f. !x. P x ==> R x (f x)) Sun 27th Oct 2013 cart.ml Added PCROSS_DIFF = |- (!s t u. s PCROSS (t DIFF u) = s PCROSS t DIFF s PCROSS u) /\ (!s t u. (s DIFF t) PCROSS u = s PCROSS u DIFF t PCROSS u) Wed 23rd Oct 2013 Makefile, pa_j_3.1x_6.11.ml [new file] Added yet another pa_j file, created by Freek Wiedijk, that appears to work correctly for all camlp5 versions 6.07 - 6.11 (current one), and modified the Makefile to select appropriately. Mon 23rd Sep 2013 nums.ml Made a modification from Tom Hales to "new_specification", simply using the names of the constants instead of some hidden counter to make the definitions unique. Fri 6th Sep 2013 cart.ml Added INTER_PCROSS = |- !s s' t t'. (s PCROSS t) INTER (s' PCROSS t') = (s INTER s') PCROSS (t INTER t') Mon 26th Aug 2013 sets.ml Added INFINITE_ENUMERATE = |- !s:num->bool. INFINITE s ==> ?r:num->num. (!m n. m < n ==> r(m) < r(n)) /\ IMAGE r (:num) = s Thu 15th Aug 2013 Library/card.ml Added CARD_LE_RELATIONAL_FULL = |- !R:A->B->bool s t. (!y. y IN t ==> ?x. x IN s /\ R x y) /\ (!x y y'. x IN s /\ y IN t /\ y' IN t /\ R x y /\ R x y' ==> y = y') ==> t <=_c s Wed 17th Jul 2013 Library/floor.ml Added INTEGER_ROUND = |- !x. ?n. integer n /\ abs(x - n) <= &1 / &2 Fri 12th Jul 2013 Library/floor.ml Added a few theorems asserting that there are integers between well-spaced pairs of reals: INTEGER_EXISTS_BETWEEN = |- !x y. x + &1 <= y ==> ?n. integer n /\ x <= n /\ n < y INTEGER_EXISTS_BETWEEN_ABS = |- !x y. &1 <= abs(x - y) ==> ?n. integer n /\ (x <= n /\ n < y \/ y <= n /\ n < x) INTEGER_EXISTS_BETWEEN_ABS_LT = |- !x y. &1 < abs(x - y) ==> ?n. integer n /\ (x < n /\ n < y \/ y < n /\ n < x) INTEGER_EXISTS_BETWEEN_ALT = |- !x y. x + &1 <= y ==> ?n. integer n /\ x < n /\ n <= y INTEGER_EXISTS_BETWEEN_LT = |- !x y. x + &1 < y ==> ?n. integer n /\ x < n /\ n < y Thu 11th Jul 2013 sets.ml Added a couple of slightly technical set lemmas: FINITE_TRANSITIVITY_CHAIN = |- !R s. FINITE s /\ (!x. ~(R x x)) /\ (!x y z. R x y /\ R y z ==> R x z) /\ (!x. x IN s ==> ?y. y IN s /\ R x y) ==> s = {} UNIONS_MAXIMAL_SETS = |- !f. FINITE f ==> UNIONS {t | t IN f /\ !u. u IN f ==> ~(t PSUBSET u)} = UNIONS f Mon 8th Jul 2013 int.ml Added a trivial but useful theorem: INT_OF_NUM_EXISTS = |- !x. (?n. x = &n) <=> &0 <= x Wed 3rd Jul 2013 real.ml, int.ml Added two simple clausal theorems for "signs of signs" REAL_SGN_INEQS = |- (!x. &0 <= real_sgn x <=> &0 <= x) /\ (!x. &0 < real_sgn x <=> &0 < x) /\ (!x. &0 >= real_sgn x <=> &0 >= x) /\ (!x. &0 > real_sgn x <=> &0 > x) /\ (!x. &0 = real_sgn x <=> &0 = x) /\ (!x. real_sgn x <= &0 <=> x <= &0) /\ (!x. real_sgn x < &0 <=> x < &0) /\ (!x. real_sgn x >= &0 <=> x >= &0) /\ (!x. real_sgn x > &0 <=> x > &0) /\ (!x. real_sgn x = &0 <=> x = &0) INT_SGN_INEQS = |- (!x. &0 <= int_sgn x <=> &0 <= x) /\ (!x. &0 < int_sgn x <=> &0 < x) /\ (!x. &0 >= int_sgn x <=> &0 >= x) /\ (!x. &0 > int_sgn x <=> &0 > x) /\ (!x. &0 = int_sgn x <=> &0 = x) /\ (!x. int_sgn x <= &0 <=> x <= &0) /\ (!x. int_sgn x < &0 <=> x < &0) /\ (!x. int_sgn x >= &0 <=> x >= &0) /\ (!x. int_sgn x > &0 <=> x > &0) /\ (!x. int_sgn x = &0 <=> x = &0) Wed 3rd Jul 2013 int.ml Fixed a bug in INT_OF_REAL_THM where theorems with more than 2 conjuncts were not being transformed correctly. Sat 29th Jun 2013 miz3/miz3.ml, miz3/bin/miz3 Made a couple of improvements to the portability of miz3 to Mac OS X, based on suggestions from Josh Jordan. Replaced the explicit signal number 12 by Sys.sigusr2, and replace realpath (which doesn't exist by default in OS X) by a line of Perl from Freek Wiedijk. Fri 28th Jun 2013 tactics.ml, Library/card.ml Moved TRANS_TAC from the cardinality theories to the main core. Fri 21st Jun 2013 Library/products.ml Added the obvious theorems about products of negations: PRODUCT_NEG = |- !f s. FINITE s ==> product s (\i. --f i) = -- &1 pow CARD s * product s f PRODUCT_NEG_NUMSEG = |- !f m n. product (m..n) (\i. --f i) = -- &1 pow ((n + 1) - m) * product (m..n) f PRODUCT_NEG_NUMSEG_1 = |- !f n. product (1..n) (\i. --f i) = -- &1 pow n * product (1..n) f Tue 4th Jun 2013 cart.ml Added a 4-element type and corresponding theorems HAS_SIZE_4 = |- (:4) HAS_SIZE 4 DIMINDEX_4 = |- dimindex (:4) = 4 Fri 31st May 2013 cart.ml Added: PCROSS_INTER = |- (!s t u. s PCROSS (t INTER u) = (s PCROSS t) INTER (s PCROSS u)) /\ (!s t u. (s INTER t) PCROSS u = (s PCROSS u) INTER (t PCROSS u)) PCROSS_UNION = |- (!s t u. s PCROSS (t UNION u) = (s PCROSS t) UNION (s PCROSS u)) /\ (!s t u. (s UNION t) PCROSS u = (s PCROSS u) UNION (t PCROSS u)) PCROSS_UNIONS_UNIONS = |- !f g. (UNIONS f) PCROSS (UNIONS g) = UNIONS {s PCROSS t | s IN f /\ t IN g} PCROSS_UNIONS = |- (!s f. s PCROSS (UNIONS f) = UNIONS {s PCROSS t | t IN f}) /\ (!f t. (UNIONS f) PCROSS t = UNIONS {s PCROSS t | s IN f}) Thu 30th May 2013 Library/card.ml Added some analogs of theorems already there for FINITE COUNTABLE_SUBSET_IMAGE = |- !f s t. COUNTABLE t /\ t SUBSET IMAGE f s <=> (?s'. COUNTABLE s' /\ s' SUBSET s /\ t = IMAGE f s') EXISTS_COUNTABLE_SUBSET_IMAGE = |- !P f s. (?t. COUNTABLE t /\ t SUBSET IMAGE f s /\ P t) <=> (?t. COUNTABLE t /\ t SUBSET s /\ P (IMAGE f t)) FORALL_COUNTABLE_SUBSET_IMAGE = |- !P f s. (!t. COUNTABLE t /\ t SUBSET IMAGE f s ==> P t) <=> (!t. COUNTABLE t /\ t SUBSET s ==> P (IMAGE f t)) Thu 30th May 2013 sets.ml Added the following FORALL_SUBSET_UNION = |- !t u:A->bool. (!s. s SUBSET t UNION u ==> P s) <=> (!t' u'. t' SUBSET t /\ u' SUBSET u ==> P(t' UNION u')) EXISTS_SUBSET_UNION = |- !t u:A->bool. (?s. s SUBSET t UNION u /\ P s) <=> (?t' u'. t' SUBSET t /\ u' SUBSET u /\ P(t' UNION u')) Sun 26th May 2013 hol.ml Added an extra line from Hendrik Tews to make things work more smoothly with OCaml 4.xx: #directory "+compiler-libs";; Fri 24th May 2013 Library/iter.ml Addded ITER_1 = |- !f x. ITER 1 f x = f x Mon 20th May 2013 Library/card.ml Added COUNTABLE_IMP_CARD_LT_REAL = |- !s. COUNTABLE s ==> s <_c (:real) Thu 2nd May 2013 sets.ml Added CARD_SING = |- !a. CARD {a} = 1 Wed 1st May 2013 cart.ml Added: PCROSS_EQ = |- !s s' t t'. s PCROSS t = s' PCROSS t' <=> (s = {} \/ t = {}) /\ (s' = {} \/ t' = {}) \/ s = s' /\ t = t' Tue 30th Apr 2013 cart.ml Added: IMAGE_FSTCART_PCROSS = |- !s t. IMAGE fstcart (s PCROSS t) = if t = {} then {} else s IMAGE_SNDCART_PCROSS = |- !s t. IMAGE sndcart (s PCROSS t) = if s = {} then {} else t Fri 5th Apr 2013 Examples/inverse_bug_puzzle_tac.ml Updated this file to a new version from Bill Richter using some slightly different constructs. Fri 5th Apr 2013 Examples/prover9.ml Fixed a bug in the prover9 interface pointed out by Keiichi Tsujimoto, where the initial preprocessing could cause the initial goal to collapse to `T` or `F`, which is then not handled properly by the main prover9 step. The fix is to treat this specially in the preprocessing phase. Sun 31st Mar 2013 Examples/inverse_bug_puzzle_miz3.ml Made an update to Bill Richter's miz3 version of the inverse bug puzzle. Sat 23rd Mar 2013 Examples/inverse_bug_puzzle_tac.ml, Examples/inverse_bug_puzzle_miz3.ml Added a tactic version Examples/inverse_bug_puzzle_tac.ml of Bill Richter's inverse bug puzzle solution, and renamed the earlier one as Examples/inverse_bug_puzzle_miz3.ml. Sat 23rd Mar 2013 Makefile Incorporated a change from Jack Pappas to ensure that the Makefile handles non-standard locations for camlp4, using "`camlp[45] -where`" instead of "+camlp[45]". Fri 22nd Mar 2013 cart.ml, Library/card.ml Added EXISTS_IN_PCROSS = |- (?z. z IN s PCROSS t /\ P z) <=> (?x y. x IN s /\ y IN t /\ P (pastecart x y)) PCROSS_MONO = |- !s t s' t'. s SUBSET s' /\ t SUBSET t' ==> s PCROSS t SUBSET s' PCROSS t' SUBSET_PCROSS;; |- !s t s' t'. s PCROSS t SUBSET s' PCROSS t' <=> s = {} \/ t = {} \/ s SUBSET s' /\ t SUBSET t' UNIV_PCROSS_UNIV = |- (:A^M) PCROSS (:A^N) = (:A^(M,N)finite_sum) FINITE_PCROSS_EQ = |- !s t. FINITE (s PCROSS t) <=> s = {} \/ t = {} \/ FINITE s /\ FINITE t COUNTABLE_CARD_MUL = |- !s t. COUNTABLE s /\ COUNTABLE t ==> COUNTABLE (s *_c t); COUNTABLE_CARD_MUL_EQ = |- !s t. COUNTABLE (s *_c t) <=> s = {} \/ t = {} \/ COUNTABLE s /\ COUNTABLE t CARD_EQ_PCROSS = |- !s t. s PCROSS t =_c s *_c t COUNTABLE_PCROSS_EQ = |- !s t. COUNTABLE (s PCROSS t) <=> s = {} \/ t = {} \/ COUNTABLE s /\ COUNTABLE t COUNTABLE_PCROSS = |- !s t. COUNTABLE s /\ COUNTABLE t ==> COUNTABLE (s PCROSS t) Thu 21st Mar 2013 cart.ml Added more theorems about PCROSS: PCROSS_EQ_EMPTY = |- !s t. s PCROSS t = {} <=> s = {} \/ t = {} HAS_SIZE_PCROSS = |- !s t m n. s HAS_SIZE m /\ t HAS_SIZE n ==> s PCROSS t HAS_SIZE m * n FINITE_PCROSS = |- !s t. FINITE s /\ FINITE t ==> FINITE (s PCROSS t) PCROSS_EMPTY = |- (!s. s PCROSS {} = {}) /\ (!t. {} PCROSS t = {}) Wed 20th Mar 2013 cart.ml Added a definition of a variant notion of product using pasting of vectors, called PCROSS, plus two elementary theorems: PCROSS = |- s PCROSS t = {pastecart (x:A^M) (y:A^N) | x IN s /\ y IN t} FORALL_IN_PCROSS = |- (!z. z IN s PCROSS t ==> P z) <=> (!x y. x IN s /\ y IN t ==> P(pastecart x y)) PASTECART_IN_PCROSS = |- !s t x y. (pastecart x y) IN (s PCROSS t) <=> x IN s /\ y IN t Mon 18th Mar 2013 preterm.ml Added a slight update from Vincent Aravantinos to the typechecking error messages, to handle additional cases. Fri 15th Mar 2013 iterate.ml Added a couple of strict positivity theorems for sums: NSUM_POS_LT = |- !f s. FINITE s /\ (?x. x IN s /\ 0 < f x) ==> 0 < nsum s f SUM_POS_LT = |- !f s. FINITE s /\ (!x. x IN s ==> &0 <= f x) /\ (?x. x IN s /\ &0 < f x) ==> &0 < sum s f Fri 15th Feb 2013 Examples/inverse_bug_puzzle.ml [new file] Added a new file due to Bill Richter giving a solution to the inverse bug puzzle in the tutorial, using miz3 and the vector theories. Fri 15th Feb 2013 tactics.ml Improved the error reporting in EXISTS_TAC, X_CHOOSE_TAC and X_GEN_TAC, which now in particular specify the expected and received types of the terms when they don't agree. Wed 13th Feb 2013 real.ml Added the following obvious theorem (which does depend on the choice of 1/0): REAL_DIV_EQ_0 = |- !x y. x / y = &0 <=> x = &0 \/ y = &0 Fri 18th Jan 2013 lists.ml, Library/prime.ml, Library/pocklington.ml, Permutation/permuted.ml Added a miscellany of theorems: APPEND_SING = |- !h t. APPEND [h] t = CONS h t MEM_APPEND_DECOMPOSE_LEFT = |- !x l. MEM x l <=> (?l1 l2. ~MEM x l1 /\ l = APPEND l1 (CONS x l2)) MEM_APPEND_DECOMPOSE = |- !x l. MEM x l <=> (?l1 l2. l = APPEND l1 (CONS x l2)) PERMUTED_APPEND_SWAP = |- !l1 l2. (APPEND l1 l2) PERMUTED (APPEND l2 l1) DIVIDES_EXP_MINUS1 = |- !k n. n - 1 divides n EXP k - 1 DIVIDES_EXP_PLUS1 = |- !n k. ODD k ==> n + 1 divides n EXP k + 1 PRIME_DIVEXP_EQ = |- !n p x. prime p ==> (p divides x EXP n <=> p divides x /\ ~(n = 0)) PRIME_POWER_EXISTS = |- !q. prime q ==> ((?i. n = q EXP i) <=> (!p. prime p /\ p divides n ==> p = q)) Thu 17th Jan 2013 Library/products.ml Added PRODUCT_CLAUSES_LEFT = |- !f m n. m <= n ==> product (m..n) f = f m * product (m + 1..n) f PRODUCT_CLAUSES_RIGHT = |- !f m n. 0 < n /\ m <= n ==> product (m..n) f = product (m..n - 1) f * f n Wed 16th Jan 2013 sets.ml Added UNIONS_MONO = |- (!x. x IN s ==> (?y. y IN t /\ x SUBSET y)) ==> UNIONS s SUBSET UNIONS t UNIONS_MONO_IMAGE = |- (!x. x IN s ==> f x SUBSET g x) ==> UNIONS (IMAGE f s) SUBSET UNIONS (IMAGE g s) Sun 6th Jan 2013 miz3/Samples/icms.ml Removed the load of Multivariate/misc.ml from this example, putting the proof of the Archimedian lemma directly in this file. There is now an incompatibility with the Multivariate use of "from" as a constant with the use as a miz3 reserved word. This should probably be dealt with more systematically, perhaps by renaming the constant to FROM. Fri 4th Jan 2013 tactics.ml, theorems.ml, help.ml, Help/DESTRUCT_TAC.doc [new file], Help/FIX_TAC.doc [new file], Help/HYP.doc [new file], Help/INTRO_TAC.doc [new file] Added four new tactic constructs from Marco Maggesi. HYP is analogous to ASM but restricting to named (labelled) assumptions. DESTRUCT_TAC, FIX_TAC and INTRO_TAC give more concise and elegant ways of fixing variables, performing introduction on a goal or elimination on a theorem, labelling assumptions in the process. Fri 4th Jan 2013 Multivariate/tarski.ml, 100/independence.ml Added two files showing Tarski's 11 axioms for geometry hold in the Euclidean plane but all except axiom 10 hold in the Klein model of the hyperbolic plane. This effectively proves the independence of the parallel postulate. Fri 4th Jan 2013 Library/floor.ml Added RATIONAL_BETWEEN = |- !a b. a < b ==> ?q. rational q /\ a < q /\ q < b Fri 4th Jan 2013 cart.ml Added PASTECART_INJ = |- !x y w z. pastecart x y = pastecart w z <=> x = w /\ y = z Wed 2nd Jan 2013 system.ml, miz3/miz3.ml Extended the regular quotation parser to pass quotations of the form `;....` to miz3's "parse_qproof". This makes it more convenient to load files with miz3 proofs in without first setting up the new parser before the file is loaded. Thu 20th Dec 2012 sets.ml Added SUBSET_INTERS = |- !s f. s SUBSET INTERS f <=> (!t. t IN f ==> s SUBSET t) Wed 19th Dec 2012 sets.ml Renamed the existing DIFF_UNIONS to UNIONS_DIFF and added the following: DIFF_UNIONS = |- !u s. u DIFF UNIONS s = u INTER INTERS {u DIFF t | t IN s} DIFF_UNIONS_NONEMPTY = |- !u s. ~(s = {}) ==> u DIFF UNIONS s = INTERS {u DIFF t | t IN s} This seems more consistent with the naming of DIFF_INTERS. Mon 17th Dec 2012 Library/wo.ml Added the definition of total order and a proof of the order extension theorem: toset = |- toset l <=> poset l /\ !x y. x IN fl(l) /\ y IN fl(l) ==> l(x,y) \/ l(y,x) OEP = |- !p. poset p ==> ?t. toset t /\ fl(t) = fl(p) /\ p SUBSET t Fri 14th Dec 2012 preterm.ml Updated the typechecker with a new version from Vincent Aravantinos that gives more useful error messages when typechecking fails. Wed 5th Dec 2012 sets.ml Added a couple of reformulations of injectivity. This is useful since the left-hand characterization can make the simplifier get very slow. INJECTIVE_ON_ALT = |- !P f. (!x y. P x /\ P y /\ f x = f y ==> x = y) <=> (!x y. P x /\ P y ==> (f x = f y <=> x = y)) INJECTIVE_ALT = |- !f. (!x y. f x = f y ==> x = y) <=> (!x y. f x = f y <=> x = y) Fri 16th Nov 2012 tactics.ml, Help/STRUCT_CASES_THEN.doc [new file] At the suggestion of Petros Papapanagiotou, made a theorem-tactic form STRUCT_CASES_THEN of STRUCT_CASES_TAC, to allow for more flexibility in how the resulting theorems get used. Fri 16th Nov 2012 IsabelleLight/*, RichterHilbertAxiomGeometry/* Brought these up to date with recent changes by their respective authors. Fri 16th Nov 2012 hol.ml, preterm.ml, parser.ml, printer.ml Reshuffled a few preliminaries between these files so that "printer.ml" could be loaded before the other two. This is a prelude to putting in code from Vincent Aravantinos for better error messages from typechecking, where it's useful to have functions from "printer.ml" for the error reporting. Fri 16th Nov 2012 sets.ml, cart.ml Added a few simple theorems about cardinalities of finite universe sets, and also a natural counterpart UNIV_GSPEC to EMPTY_GSPEC. New theorems: CARD_BOOL = |- CARD (:bool) = 2 CARD_CART_UNIV = |- FINITE (:A) ==> CARD (:A^N) = CARD (:A) EXP dimindex (:N) CARD_FUNSPACE_UNIV = |- FINITE (:A) /\ FINITE (:B) ==> CARD (:A->B) = CARD (:B) EXP CARD (:A) FINITE_BOOL = |- FINITE (:bool) FINITE_CART_UNIV = |- FINITE (:A) ==> FINITE (:A^N) FINITE_FUNSPACE_UNIV = |- FINITE (:A) /\ FINITE (:B) ==> FINITE (:A->B) HAS_SIZE_BOOL = |- (:bool) HAS_SIZE 2 HAS_SIZE_CART_UNIV = |- !m. (:A) HAS_SIZE m ==> (:A^N) HAS_SIZE m EXP dimindex (:N) HAS_SIZE_FUNSPACE_UNIV = |- !m n. (:A) HAS_SIZE m /\ (:B) HAS_SIZE n ==> (:A->B) HAS_SIZE n EXP m UNIV_GSPEC = |- {x | T} = UNIV Tue 13th Nov 2012 sets.ml, Library/card.ml Added three new theorems about COUNTABLE in exact analogy with FINITE, and also cleaned up the FINITE proof of one of them based on the pattern. COUNTABLE_IMAGE_INJ = |- !f A. (!x y. f x = f y ==> x = y) /\ COUNTABLE A ==> COUNTABLE {x | f x IN A} COUNTABLE_IMAGE_INJ_EQ = |- !f s. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> (COUNTABLE (IMAGE f s) <=> COUNTABLE s) COUNTABLE_IMAGE_INJ_GENERAL = |- !f A s. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) /\ COUNTABLE A ==> COUNTABLE {x | x IN s /\ f x IN A} Fri 28th Sep 2012 pair.ml, int.ml, sets.ml, database.ml, Library/card.ml Added a few miscellaneous little theorems: CHOOSE_SUBSET_BETWEEN = |- !n s u. s SUBSET u /\ FINITE s /\ CARD s <= n /\ (FINITE u ==> n <= CARD u) ==> (?t. s SUBSET t /\ t SUBSET u /\ t HAS_SIZE n) EXISTS_CURRY = |- !P. (?f. P f) <=> (?f. P (\(a,b). f a b)) FORALL_CURRY = |- !P. (!f. P f) <=> (!f. P (\(a,b). f a b)) FORALL_FINITE_SUBSET_IMAGE = |- !P f s. (!t. FINITE t /\ t SUBSET IMAGE f s ==> P t) <=> (!t. FINITE t /\ t SUBSET s ==> P (IMAGE f t)) INFINITE_SUPERSET = |- !s t. INFINITE s /\ s SUBSET t ==> INFINITE t INF_SING = |- !a. inf {a} = a PAIRED_ETA_THM = |- (!f. (\(x,y). f (x,y)) = f) /\ (!f. (\(x,y,z). f (x,y,z)) = f) /\ (!f. (\(w,x,y,z). f (w,x,y,z)) = f) SUP_SING = |- !a. sup {a} = a as well as slightly reshuffling some existing theorems and adding new cardinality theorems. Thu 6th Sep 2012 update_database.ml Fixed the update_database code for OCaml 4.00. This entails making a couple more things conditional on the version number (using the exec "code" method already used in Roland Zumkeller's original code). First, the "Tvar" constructor now takes an argument (and because of the way OCaml's internal representation works with different sequences for nullary and non-nullary, this changes some of the constructor numbers). Second, the tbl type has changed to the use of EnvTbl that effectively wraps up the core type 'a into 'a * bool ref. Fri 31st Aug 2012 Makefile, README Fixed up the Makefile both for recent camlp5s (not exhaustively tested, but I think it is right) and also for OCaml 4.00, as well as making the instructions in the README file a bit more up-to-date. Fri 1st Jun 2012 100/descartes.ml [new file] Added a proof of Descartes's rule of signs, based on Rob Arthan's paper "Descartes's Rule of Signs by an Easy Induction". Wed 30th May 2012 100/cayley_hamilton.ml [new file] Added a proof of the Cayley-Hamilton theorem for real matrices. Wed 30th May 2012 100/feuerbach.ml [new file] Added a proof of Feuerbach's theorem. Tue 29th May 2012 Examples/sylvester_gallai.ml [new file] Added a proof of the Sylvester-Gallai theorem. Mon 28th May 2012 100/morley.ml [new file] Added a proof of Morley's theorem, following Alain Connes's paper "A new proof of Morley's theorem". Sun 27th May 2012 Examples/brunn_minkowski.ml [new file] Added a proof of the Brunn-Minkowski theorem. Sat 26th May 2012 100/platonic.ml [new file] Added a proof that the Platonic solids are limited to the classic five, and that those all do exist and are regular. Because of the rather crude way in which the computations of facial structure are done, this proof takes quite a long time to load. Sat 26th May 2012 QBF/* [new files] Added Ondrej Kuncar's code for constructing proofs of quantified boolean formulas in HOL Light using proof traces from Squolem. Fri 25th May 2012 100/polyhedron.ml [new file] Added a proof of Euler's polyhedron formula for convex polytopes, and indeed the general Euler-Poincare relation. This follows Jim Lawrence's "A short proof of Euler's relation for convex polytopes" (Canadian Math Bulletin, 1997). This proof was mostly quite easy and natural to formalize, except for some slightly tedious switching between open and closed cells at one point. Wed 16th May 2012 Makefile Added a test to the Makefile to catch camlp5 6.05 and use the pa_j.ml designed for 6.02.2, which according to a test by Bill Richter works fine. I could probably also add the intermediate versions 6.03 and 6.04, but I didn't do so for now. Wed 16th May 2012 miz3/miz3.ml, miz3/test.ml, holtest Made a small tweak to miz3.ml from Freek, and removed the test run in the final line. Also replaced the former test.ml file with the former Samples/ALL (which is no longer there now), and made "holtest" run it twice in case cacheing changes things. Mon 7th May 2012 sets.ml Added the trivial but sometimes useful lemma INFINITE_SUPERSET = |- !s t. INFINITE s /\ s SUBSET t ==> INFINITE t Fri 4th May 2012 miz3/* [new files], passim Added Freek Wiedijk's miz3 mode, and since there I first made a tweak of open_in -> Pervasives.open_in (to avoid possible namespace conflicts with Multivariate/topology.ml) made a similar change in a few other places too. Fri 27th Apr 2012 sets.ml, Help/new_inductive_set.doc [new file] Added a function from Marco Maggesi for defining sets inductively by analogy with new_inductive_definition for relations. Fri 27th Apr 2012 preterm.ml, Help/the_implicit_types.doc [new file], Help/passim Added another extra feature to term parsing from Marco Maggesi, a list of implicit type schemes for variables in quotations, "the_implicit_types". Thu 26th Apr 2012 preterm.ml, Help/type_invention_error.doc [new file], type_invention_warning.doc Added code from Marco Maggesi to incorporate an additional option "type_invention_error" that forces type variables to be an error, not merely be warned about. Added a corresponding documentation file and some cross-referencing w.r.t the existing file "type_invention_warning.doc". Thu 26th Apr 2012 holtest Incorporated some changes from Hendrik Tews to make "holtest" a bit more robust and keep it consistent with the HOL Light debian package. Tue 24th Apr 2012 Help/new_axiom.doc Made a couple of small fixes to the "new_axiom" documentation, removing a stray paren and adding a cross-reference to "axioms". Fri 30th Mar 2012 Permutation/permuted.ml Added a couple of lemmas about list permutations preserving "PAIRWISE" for a symmetric relation, plus strong induction PERMUTED_INDUCT_STRONG: PERMUTED_IMP_PAIRWISE = |- !P l l'. (!x y. P x y ==> P y x) /\ l PERMUTED l' /\ PAIRWISE P l ==> PAIRWISE P l' PERMUTED_PAIRWISE = |- !P l l'. (!x y. P x y ==> P y x) /\ l PERMUTED l' ==> (PAIRWISE P l <=> PAIRWISE P l') Thu 29th Mar 2012 Library/card.ml Added a few new theorems about cardinal arithmetic: CARD_LE_RELATIONAL = |- !R. (!x y y'. x IN s /\ R x y /\ R x y' ==> y = y') ==> {y | ?x. x IN s /\ R x y} <=_c s CARD_LT_FINITE_INFINITE = |- !s t. FINITE s /\ INFINITE t ==> s <_c t CARD_ADD2_ABSORB_LT = |- !s t u. INFINITE u /\ s <_c u /\ t <_c u ==> s +_c t <_c u CARD_ADD_FINITE_EQ = |- !s t. FINITE (s +_c t) <=> FINITE s /\ FINITE t CARD_ADD_C = |- !s t. FINITE s /\ FINITE t ==> CARD (s +_c t) = CARD s + CARD t CARD_LT_ADD = |- !s s' t t'. s <_c s' /\ t <_c t' ==> s +_c t <_c s' +_c t' Fri 9th Mar 2012 Library/card.ml Added a few more trivial lemmas about cardinality: CARD_LT_LE = |- !s t. s <_c t <=> s <=_c t /\ ~(s =_c t) CARD_LE_LT = |- !s t. s <=_c t <=> s <_c t \/ s =_c t COUNTABLE_ALT = |- !s. COUNTABLE s <=> s <=_c (:num) COUNTABLE_CASES = |- !s. COUNTABLE s <=> FINITE s \/ s =_c (:num) Sun 26th Feb 2012 sets.ml Added CARD_IMAGE_EQ_INJ = |- !f:A->B s. FINITE s ==> (CARD(IMAGE f s) = CARD s <=> !x y. x IN s /\ y IN s /\ f x = f y ==> x = y) PAIRWISE_IMAGE = |- !r f. pairwise r (IMAGE f s) <=> pairwise (\x y. ~(f x = f y) ==> r (f x) (f y)) s Thu 23rd Feb 2012 sets.ml Added PAIRWISE_INSERT = |- !r x s. pairwise r (x INSERT s) <=> (!y. y IN s /\ ~(y = x) ==> r x y /\ r y x) /\ pairwise r s Tue 21st Feb 2012 sets.ml, Library/card.ml Added two direct equality results for sups and infs: SUP_EQ = |- !s t. (!b. (!x. x IN s ==> x <= b) <=> (!x. x IN t ==> x <= b)) ==> sup s = sup t INF_EQ = |- !s t. (!a. (!x. x IN s ==> a <= x) <=> (!x. x IN t ==> a <= x)) ==> inf s = inf t and a few basic lemmas about cardinality: CARD_LE_COUNTABLE = |- !s t. COUNTABLE t /\ s <=_c t ==> COUNTABLE s CARD_COUNTABLE_CONG = |- !s t. s =_c t ==> (COUNTABLE s <=> COUNTABLE t) CARD_EQ_FINITE = |- !s t. FINITE t /\ s =_c t ==> FINITE s CARD_EQ_COUNTABLE = |- !s t. COUNTABLE t /\ s =_c t ==> COUNTABLE s CARD_EQ_IMAGE = |- !f s. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> IMAGE f s =_c s CARD_EQ_REAL_IMP_UNCOUNTABLE = |- !s. s =_c (:real) ==> ~COUNTABLE s Sat 11th Feb 2012 Library/floor.ml Added the obvious facts that the integer and rational valued reals are both infinite: INFINITE_INTEGER = |- INFINITE integer INFINITE_RATIONAL = |- INFINITE rational Thu 2nd Feb 2012 iterate.ml Added two theorems about pushing MODs through sums of natural numbers: MOD_NSUM_MOD = |- !f n s. FINITE s /\ ~(n = 0) ==> nsum s f MOD n = nsum s (\i. f i MOD n) MOD n MOD_NSUM_MOD_NUMSEG = |- !f a b n. ~(n = 0) ==> nsum (a..b) f MOD n = nsum (a..b) (\i. f i MOD n) MOD n Wed 1st Feb 2012 arith.ml Added a theorem that's a trivial consequence of DIVISION, but sometimes a bit more convenient DIVISION_SIMP = |- (!m n. ~(n = 0) ==> m DIV n * n + m MOD n = m) /\ (!m n. ~(n = 0) ==> n * m DIV n + m MOD n = m) Wed 1st Feb 2012 Mizarlight/Makefile Fixed up the Makefile to choose camlp5 versus camlp4 in a correct way as in the main system Makefile (previously this failed for OCaml >= 3.12). Tue 31st Jan 2012 sets.ml Added one more set lemma, analogous to INTER_UNIONS: DIFF_UNIONS = |- !s t. (UNIONS s) DIFF t = UNIONS {x DIFF t | x IN s} Thu 26th Jan 2012 cart.ml Added a somewhat technical lemma, a generalization of FINITE_INDEX_INRANGE that is useful for obscure situations to get rid of range conditions in doubly indexed cartesian products: FINITE_INDEX_INRANGE_2 = |- !i. ?k. 1 <= k /\ k <= dimindex(:N) /\ (!x:A^N. x$i = x$k) /\ (!y:B^N. y$i = y$k) Thu 19th Jan 2012 Library/card.ml Generalized COUNTABLE_PRODUCT_DEPENDENT to have the same form as FINITE_PRODUCT_DEPENDENT instead of being tied specifically to pairs. Sat 14th Jan 2012 sets.ml Added "localized" forms of the theorems about factoring functions through each other: FUNCTION_FACTORS_LEFT_GEN = |- !P f g. (!x y. P x /\ P y /\ g x = g y ==> f x = f y) <=> (?h. !x. P x ==> f(x) = h(g x)) FUNCTION_FACTORS_RIGHT_GEN = |- !P f g. (!x. P x ==> ?y. g(y) = f(x)) <=> (?h. !x. P x ==> f(x) = g(h x)) Mon 9th Jan 2012 list.ml Added two new list theorems MAP_REVERSE = |- !f l. REVERSE(MAP f l) = MAP f (REVERSE l) ALL_FILTER = |- !P Q l. ALL P (FILTER Q l) <=> ALL (\x. Q x ==> P x) l Thu 23rd Dec 2011 Library/card.ml Added CARD_LE_UNIV = |- !s:A->bool. s <=_c (:A) Wed 22nd Dec 2011 sets.ml, iterate.ml, define.ml, Library/permutations.ml, Library/products.ml, Library/iter.ml, Multivariate/ passim Did a bit of cleaning up of theorems to remove redundant quantifiers (put there by accident for variables that aren't free in the body). This includes the following and numerous theorems in the Multivariate theories: ========= sets.ml ============ SIMPLE_IMAGE_GEN HAS_SIZE_0 ========= iterate.ml ============ NSUM_BOUND_GEN NSUM_BOUND_LT_GEN NSUM_EQ_0_NUMSEG SUM_BOUND_GEN SUM_BOUND_LT_GEN SUM_EQ_0_NUMSEG ========= define.ml ============ ADMISSIBLE_MATCH ========= Library/permutations.ml ============ PERMUTES_COMPOSE ========= Library/products.ml ============ PRODUCT_EQ_1_NUMSEG ========= Library/iter.ml ============ ITER_ALT_POINTLESS Sun 18th Dec 2011 cart.ml Added a convenient set elimination theorem for pasted vectors: IN_ELIM_PASTECART_THM = |- !P a b. pastecart a b IN {pastecart x y | P x y} <=> P a b Thu 1st Dec 2011 Library/card.ml Added a few more theorems on cardinality of list and cartesian product types and also the cardinality and uncountability of the reals. This last fact is done in a primitive way to avoid depending on the analytical theories. The theorem UNCOUNTABLE_REAL has been taken out of the Multivariate theories now in consequence. CARD_EQ_LIST = |- INFINITE (:A) ==> (:(A)list) =_c (:A) CARD_EQ_CART = |- INFINITE (:A) ==> (:A^N) =_c (:A) CARD_EQ_REAL = |- (:real) =_c (:num->bool) UNCOUNTABLE_REAL = |- ~COUNTABLE (:real) Thu 1st Dec 2011 sets.ml Moved all the definitions and basic theorems about sup and inf to the end of this file from Multivariate/misc.ml: inf = |- !s. inf s = (@a. (!x. x IN s ==> a <= x) /\ (!b. (!x. x IN s ==> b <= x) ==> b <= a)) sup = |- !s. sup s = (@a. (!x. x IN s ==> x <= a) /\ (!b. (!x. x IN s ==> x <= b) ==> a <= b)))] INF = |- !s. ~(s = {}) /\ (?b. !x. x IN s ==> b <= x) ==> (!x. x IN s ==> inf s <= x) /\ (!b. (!x. x IN s ==> b <= x) ==> b <= inf s) INF_FINITE = |- !s. FINITE s /\ ~(s = {}) ==> inf s IN s /\ (!x. x IN s ==> inf s <= x) INF_FINITE_LEMMA = |- !s. FINITE s /\ ~(s = {}) ==> (?b. b IN s /\ (!x. x IN s ==> b <= x)) INF_INSERT_FINITE = |- !x s. FINITE s ==> inf (x INSERT s) = (if s = {} then x else min x (inf s)) INF_UNIQUE_FINITE = |- !s. FINITE s /\ ~(s = {}) ==> (inf s = a <=> a IN s /\ (!y. y IN s ==> a <= y)) REAL_ABS_INF_LE = |- !s a. ~(s = {}) /\ (!x. x IN s ==> abs x <= a) ==> abs (inf s) <= a REAL_ABS_SUP_LE = |- !s a. ~(s = {}) /\ (!x. x IN s ==> abs x <= a) ==> abs (sup s) <= a REAL_INF_ASCLOSE = |- !s l e. ~(s = {}) /\ (!x. x IN s ==> abs (x - l) <= e) ==> abs (inf s - l) <= e REAL_INF_BOUNDS = |- !s a b. ~(s = {}) /\ (!x. x IN s ==> a <= x /\ x <= b) ==> a <= inf s /\ inf s <= b REAL_INF_LE_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (inf s <= a <=> (?x. x IN s /\ x <= a)) REAL_INF_LT_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (inf s < a <=> (?x. x IN s /\ x < a)) REAL_INF_UNIQUE = |- !s b. (!x. x IN s ==> b <= x) /\ (!b'. b < b' ==> (?x. x IN s /\ x < b')) ==> inf s = b REAL_LE_INF = |- !b. ~(s = {}) /\ (!x. x IN s ==> b <= x) ==> b <= inf s REAL_LE_INF_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (a <= inf s <=> (!x. x IN s ==> a <= x)) REAL_LE_INF_SUBSET = |- !s t. ~(t = {}) /\ t SUBSET s /\ (?b. !x. x IN s ==> b <= x) ==> inf s <= inf t REAL_LE_SUP_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (a <= sup s <=> (?x. x IN s /\ a <= x)) REAL_LT_INF_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (a < inf s <=> (!x. x IN s ==> a < x)) REAL_LT_SUP_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (a < sup s <=> (?x. x IN s /\ a < x)) REAL_SUP_ASCLOSE = |- !s l e. ~(s = {}) /\ (!x. x IN s ==> abs (x - l) <= e) ==> abs (sup s - l) <= e REAL_SUP_BOUNDS = |- !s a b. ~(s = {}) /\ (!x. x IN s ==> a <= x /\ x <= b) ==> a <= sup s /\ sup s <= b REAL_SUP_EQ_INF = |- !s. ~(s = {}) /\ (?B. !x. x IN s ==> abs x <= B) ==> (sup s = inf s <=> (?a. s = {a})) REAL_SUP_LE = |- !b. ~(s = {}) /\ (!x. x IN s ==> x <= b) ==> sup s <= b REAL_SUP_LE_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (sup s <= a <=> (!x. x IN s ==> x <= a)) REAL_SUP_LE_SUBSET = |- !s t. ~(s = {}) /\ s SUBSET t /\ (?b. !x. x IN t ==> x <= b) ==> sup s <= sup t REAL_SUP_LT_FINITE = |- !s a. FINITE s /\ ~(s = {}) ==> (sup s < a <=> (!x. x IN s ==> x < a)) REAL_SUP_UNIQUE = |- !s b. (!x. x IN s ==> x <= b) /\ (!b'. b' < b ==> (?x. x IN s /\ b' < x)) ==> sup s = b SUP = |- !s. ~(s = {}) /\ (?b. !x. x IN s ==> x <= b) ==> (!x. x IN s ==> x <= sup s) /\ (!b. (!x. x IN s ==> x <= b) ==> sup s <= b) SUP_FINITE = |- !s. FINITE s /\ ~(s = {}) ==> sup s IN s /\ (!x. x IN s ==> x <= sup s) SUP_FINITE_LEMMA = |- !s. FINITE s /\ ~(s = {}) ==> (?b. b IN s /\ (!x. x IN s ==> x <= b)) SUP_INSERT_FINITE = |- !x s. FINITE s ==> sup (x INSERT s) = (if s = {} then x else max x (sup s)) SUP_UNIQUE_FINITE = |- !s. FINITE s /\ ~(s = {}) ==> (sup s = a <=> a IN s /\ (!y. y IN s ==> y <= a)) Thu 1st Dec 2011 real.ml Moved three basic Archimedian properties back to here from Multivariate/misc.ml: REAL_ARCH_SIMPLE = |- !x. ?n. x <= &n REAL_ARCH_LT = |- !x. ?n. x < &n REAL_ARCH = |- !x. &0 < x ==> (!y. ?n. y < &n * x) Thu 1st Dec 2011 list.ml Added EL_MAP = |- !f n l. n < LENGTH l ==> EL n (MAP f l) = f (EL n l) Mon 21st Nov 2011 iterate.ml Added a few lemmas about the behaviour of real polynomial functions, since these are quite often useful: REAL_SUB_POLYFUN = |- !a x y n. 1 <= n ==> sum(0..n) (\i. a i * x pow i) - sum(0..n) (\i. a i * y pow i) = (x - y) * sum(0..n - 1) (\j. sum(j + 1..n) (\i. a i * y pow (i - j - 1)) * x pow j) REAL_SUB_POLYFUN_ALT = |- !a x y n. 1 <= n ==> sum(0..n) (\i. a i * x pow i) - sum(0..n) (\i. a i * y pow i) = (x - y) * sum(0..n - 1) (\j. sum(0..n - j - 1) (\k. a (j + k + 1) * y pow k) * x pow j) REAL_POLYFUN_ROOTBOUND = |- !n c. ~(!i. i IN 0..n ==> c i = &0) ==> FINITE {x | sum(0..n) (\i. c i * x pow i) = &0} /\ CARD {x | sum(0..n) (\i. c i * x pow i) = &0} <= n REAL_POLYFUN_FINITE_ROOTS = |- !n c. FINITE {x | sum(0..n) (\i. c i * x pow i) = &0} <=> (?i. i IN 0..n /\ ~(c i = &0)) REAL_POLYFUN_EQ_0 = |- !n c. (!x. sum(0..n) (\i. c i * x pow i) = &0) <=> (!i. i IN 0..n ==> c i = &0) REAL_POLYFUN_EQ_CONST = |- !n c k. (!x. sum(0..n) (\i. c i * x pow i) = k) <=> c 0 = k /\ (!i. i IN 1..n ==> c i = &0) Thu 3rd Nov 2011 arith.ml Added another somewhat intricate DIV/MOD theorem: MOD_MOD_EXP_MIN = |- !x p m n. ~(p = 0) ==> x MOD p EXP m MOD p EXP n = x MOD p EXP MIN m n Tue 1st Nov 2011 arith.ml Added the following, a natural counterpart to MOD_MULT_ADD: DIV_MULT_ADD = |- !a b n. ~(n = 0) ==> (a * n + b) DIV n = a + b DIV n Thu 20th Oct 2011 hol.ml, class.ml, nums.ml, arith.ml, recursion.ml, pair.ml Based on an idea discussed with Mark Adams, changed "new_specification" to exploit dummy quantifiers and numeral tags to avoid provable equalities between different constants introduced with new_specification based on the same (or a provably equivalent) existence theorem. Since this requires the existence of numerals, quite a bit of build order reshuffling was necessary, and two former uses of new_specification (IND_SUC/IND_0 and BIT0_DEF) are now done "manually". Fri 7th Oct 2011 Makefile Made a fix to the Makefile from Ramana Kumar on the hol-light page (see issue 2). This uses the file pa_j_3.1x_6.02.2.ml for version 6.03 of OCaml as well. I just tried it with Ocaml 3.12.1 and camlp5 6.02.3, and it seems to work fine. Mon 29th Aug 2011 calc_rat.ml Changed REAL_RAT_DIV_CONV not to fail if it is given a canonical rational; instead used CHANGED_CONV in the depth conversion in REAL_RAT_REDUCE_CONV. This is consistent with the use of REAL_RAT_NEG_CONV, makes REAL_RAT_DIV_CONV a more convenient building block, and agrees with the current documentation. Wed 24th Aug 2011 real.ml, int.ml Added theorems about the real and integer sign function: REAL_SGN_CASES = |- !x. real_sgn x = &0 \/ real_sgn x = &1 \/ real_sgn x = -- &1 REAL_SGN_EQ = |- (!x. real_sgn x = &0 <=> x = &0) /\ (!x. real_sgn x = &1 <=> x > &0) /\ (!x. real_sgn x = -- &1 <=> x < &0) with INT_SGN_CASES and INT_SGN_EQ being the same for integers. Wed 24th Aug 2011 sets.ml Added two natural theorems about intersections of unions INTERS_UNION = |- !s t. INTERS (s UNION t) = INTERS s INTER INTERS t INTERS_OVER_UNIONS;; |- !f s. INTERS {UNIONS (f x) | x IN s} = UNIONS {INTERS {g x | x IN s} | g | !x. x IN s ==> g x IN f x} Thu 18th Aug 2011 100/inclusion_exclusion.ml Added versions of inclusion-exclusion for additive real functions (INCLUSION_EXCLUSION_REAL_RESTRICTED_INDEXED, INCLUSION_EXCLUSION_REAL_RESTRICTED, INCLUSION_EXCLUSION_REAL_INDEXED and INCLUSION_EXCLUSION_REAL), deriving the cardinality version from those. Renamed the old INCLUSION_EXCLUSION_REAL as INCLUSION_EXCLUSION_REAL_FUNCTION and commented it out together with its dependency on Library/products.ml Sun 31st Jul 2011 IsabelleLight/* [new files] Installed "Isabelle Light" from Petros Papapanagiotou and Jacques Fleuriot, which consists of implementations of Isabelle-style tactics and other user-friendly shortcuts. Sun 10th Jul 2011 100/fourier.ml [new file] Added a file with some material on L_p spaces and Fourier series. Tue 28th Jun 2011 pair.ml, parser.ml Fixed a LET_TAC bug pointed out by Vu Khac Ky: failures were occurring if the same variable was used on the left and right of a pattern (except if they are a trivial let, i.e. corresponding lefts and rights exactly equal), e.g. g `v IN V DIFF (let k,ul = (P:num#(real^3)list->num#(real^3)list) (k,ul) in (Q:num#(real^3)list->real^3->bool) (k,ul)) ==> F`;; e LET_TAC;; Now the abbreviating variables are renamed as necessary to avoid clashes with free variables of the goal or each other (some of this would previously happen later in the ABBREV_TAC step anyway in some cases, but not this one). Looking at such corner cases made me add a parsing warning if there are multiple uses of the same variable in different arms of a simultaneous let binding, e.g. `let x = 1 and x = 2 in x = 42`, in which case the first one is hidden. Thu 9th Jun 2011 realarith.ml Made a change in the variable-choosing heuristic in Fourier-Motzkin, motivated by the disastrous slowness of this trivial example: time REAL_ARITH `c < e / &3 /\ i1 < e / &12 /\ i3 < e / &12 /\ i2 < e / &12 /\ i4 < e / &12 ==> (i1 + i2) + (c + i3 + i4) <= &2 * e / &3`;; The previous heuristic chose a variable that minimized the number of inequalities after elimination, which is p * n + z where p, n and z are the number of current inequalities where that variable occurs positively, negatively and not at all respectively. The downside is that it doesn't take any account of the complexity of the resulting inequalities, so in a prototypical example like the above x1 < e /\ .... /\ xn < e ==> x1 + ... + xn < n * e all variables give equal blowup of n, and so e may get chosen from among the available alternatives. That's something of a disaster since from that point on you get doubly exponential blowup per stage. The new heuristic is simply p * n, counting the number of new inequalities created by the elimination step. It makes this sort of example, which is probably pretty common, dramatically better. It also seems to lead to a small but notable average speedup. Mon 23rd May 2011 Library/floor.ml, Library/sets.ml Added COUNTABLE_RESTRICT = |- !s P. COUNTABLE s ==> COUNTABLE {x | x IN s /\ P x} RATIONAL_APPROXIMATION_STRADDLE = |- !x e. &0 < e ==> ?a b. rational a /\ rational b /\ a < x /\ x < b /\ abs(b - a) < e Sun 22nd May 2011 sets.ml Added the following duality for unions and intersections: DIFF_INTERS = |- !u s. u DIFF INTERS s = UNIONS {u DIFF t | t IN s} INTERS_UNIONS = |- !s. INTERS s = UNIV DIFF (UNIONS {UNIV DIFF t | t IN s}) UNIONS_INTERS = |- !s. UNIONS s = UNIV DIFF (INTERS {UNIV DIFF t | t IN s}) Thu 12th May 2011 Library/floor.ml FRAC_UNIQUE = |- !x a. integer(x - a) /\ &0 <= a /\ a < &1 <=> frac x = a REAL_FRAC_EQ = |- !x. frac x = x <=> &0 <= x /\ x < &1 Tue 5th Apr 2011 Library/floor.ml REAL_FLOOR_LE = |- !x n. integer n ==> (floor x <= n <=> x - &1 < n) HAS_SIZE_INTSEG_INT = |- !a b. integer a /\ integer b ==> {x | integer x /\ a <= x /\ x <= b} HAS_SIZE (if b < a then 0 else num_of_int(int_of_real(b - a + &1))) CARD_INTSEG_INT = |- !a b. integer a /\ integer b ==> CARD {x | integer x /\ a <= x /\ x <= b} = (if b < a then 0 else num_of_int(int_of_real(b - a + &1))) REAL_CARD_INTSEG_INT = |- !a b. integer a /\ integer b ==> &(CARD {x | integer x /\ a <= x /\ x <= b}) = (if b < a then &0 else b - a + &1) Tue 5th Apr 2011 sets.ml Added POWERSET_CLAUSES = |- {s | s SUBSET {}} = {{}} /\ (!a t. {s | s SUBSET a INSERT t} = {s | s SUBSET t} UNION IMAGE (\s. a INSERT s) {s | s SUBSET t}) Fri 1st Apr 2011 Makefile, pa_j_3.1x_6.02.2.ml [new file] Once again there is an incompatibility with the latest version of camlp5, this time 6.02.2, as was reported by Kevin S. Van Horn. Made yet another appropriate version of pa_j.ml, and put another case split in the Makefile. Fri 1st Apr 2011 Minisat/minisat_parse.ml Rewrote readTrace_aux in a slightly different style, avoiding wrapping an exception handler round the recursive calls, since this had the effect of making it not be tail-recursive. Thu 31st Mar 2011 sets.ml Added the following IMAGE_INJECTIVE_IMAGE_OF_SUBSET = |- !f s. ?t. t SUBSET s /\ IMAGE f s = IMAGE f t /\ (!x y. x IN t /\ y IN t /\ f x = f y ==> x = y) Tue 22nd Mar 2011 sets.ml Added the following, a natural dual of EXISTS_SUBSET_IMAGE: FORALL_SUBSET_IMAGE = |- !P f s. (!t. t SUBSET IMAGE f s ==> P t) <=> (!t. t SUBSET s ==> P (IMAGE f t)) Tue 15th Mar 2011 fusion.ml Made a small but significant change to the kernel, based on an observation by Ondrej Kuncar, optimizing the pointer-EQ shortcut in "orda" so that it does not require an empty environment but merely one full of identical pairs. This can substantially speed up many derived rules by allowing efficient use of proformas even inside bound variables. Sat 5th Mar 2011 bool.ml Noticed that GEN is not constant-time, which is bad. On closer inspection the problem was that the abstraction that's used to instantiate the proforma is created separately from the abstracted input theorem, so the pointer-EQ subterms are hidden under an abstraction. As well as fixing that, made it more efficient in two other ways: used EQ_MP instead of PROVE_HYP, and added partial evaluation based on the first argument. Fri 25th Feb 2011 Library/poly.ml Fixed an error where override_interface was used for "divides", hiding all the other desired overloadings. Also added many new polynomial theorems from Jesse Bingham's e-transcendence proof: HD_POLY_ADD HD_POLY_CMUL HD_POLY_EXP HD_POLY_EXP_X_SUC HD_POLY_MUL HD_POLY_MUL_X ITERATE_RADD_POLYADD MONOIDAL_POLY_ADD NOT_POLY_CMUL_NIL NOT_POLY_EXP_NIL NOT_POLY_EXP_X_NIL NOT_POLY_MUL_NIL POLYDIFF_ADD POLY_ADD_ASSOC POLY_ADD_IDENT POLY_ADD_LENGTH POLY_ADD_NEUTRAL POLY_ADD_SYM POLY_CMUL_LENGTH POLY_CMUL_LID POLY_CMUL_POLY_DIFF POLY_DIFF_AUX_ADD_LEMMA POLY_DIFF_AUX_POLY_CMUL POLY_EXP_X_LENGTH POLY_EXP_X_REC POLY_MUL_LENGTH POLY_MUL_LENGTH2 POLY_MUL_LID POLY_MUL_RID POLY_SUM_EQUIV TL_POLY_CMUL TL_POLY_EXP_X_SUC TL_POLY_MUL_X Fri 25th Feb 2011 fusion.ml, ind_defs.ml, simp.ml, define.ml, Library/analysis.ml, Library/transc.ml Removed a few unused variables in these files, after noticing OCaml warnings about them. It makes things minutely smaller and faster, I suppose. Tue 22nd Feb 2011 lists.ml Added two additional list lemmas taken from Jesse Bingham's e-transcendence proof, which seem to be quite generally useful. ALL_EL = |- !P l. (!i. i < LENGTH l ==> P (EL i l)) <=> ALL P l CONS_HD_TL = |- !l. ~(l = []) ==> l = CONS (HD l) (TL l) Thu 17th Feb 2011 Examples/sos.ml Renamed a few of the ML operations on vectors since the existing names can lead to annoying name clashes with the names of theorems in the Multivariate/vectors.ml theories. vector_0 -> vec_0 dim -> vec_dim vector_const -> vec_const vector_1 -> vec_1 vector_cmul -> vec_cmul vector_neg -> vec_neg vector_add -> vec_add vector_sub -> vec_sub vector_dot -> vec_dot vector_of_list -> vec_of_list Of course, it is arguable that all this should be hidden anyway, but since this file is still somewhat experimental I don't want to tidy up and modularize the namespace just yet. Mon 14th Feb 2011 int.ml, Library/floor.ml Removed the constant "is_int" and replaced it with "integer", moving the definition of that back from the "Library/floor.ml" file to the "int.ml" file. Kept the theorem "is_int" for compatibility, with just the different constant, and removed the now redundant/meaningless INTEGER_IS_INT. Tue 8th Feb 2011 Makefile, pa_j_3.1x_6.02.1.ml [new file] Added (directly from Ondrej Kuncar) an updated version of pa_j for the latest sub-version of camlp5, and modified the Makefile to use it as needed (though this selection logic is starting to get ugly). Tue 21st Dec 2010 real.ml, database.ml Added a lemma from Valentina Bruno used in the Multivariate complex analysis theories, plus other variants. REAL_LT_LINV = |- !x y. &0 < y /\ inv y < x ==> inv x < y REAL_LT_RINV = |- !x y. &0 < x /\ x < inv y ==> y < inv x REAL_LE_LINV = |- !x y. &0 < y /\ inv y <= x ==> inv x <= y REAL_LE_RINV = |- !x y. &0 < x /\ x <= inv y ==> y <= inv x Wed 15th Dec 2010 system.ml At the suggestion of Ondrej Kuncar, switched from fun set_jrh_lexer -> set_jrh_lexer;; to just set_jrh_lexer;; This magic variable gets mapped to a constant anyway, so there is no need for this elaborate expression, which moreover generates a warning about a non-exhaustive match. Mon 13th Dec 2010 sets.ml Added SET_PAIR_THM = |- !P. {p | P p} = {a,b | P (a,b)} Thu 2nd Dec 2010 Makefile, Proofrecording/hol_light/Makefile, pa_j_3.1x_5.xx.ml [new file], pa_j_3.1x_6.xx.ml [new file], pa_j_3.10.ml [deleted], pa_j_3.11.ml [deleted] Ondrej Kuncar pointed out that the system doesn't build with the new camlp5 6.0. So now for OCaml >= 3.10 I need to start distinguishing only based on the camlp5 version. Modified the Makefile accordingly, making the two current versions for < 6.0 and >= 6.0 of camlp5 called pa_j_3.1x_5.xx.ml (what used to be pa_j_3.10.ml and pa_j_3.11.ml) and pa_j_3.1x_6.xx.ml (a new one derived by modifying the source to camlp5 6.02.0). Thu 18th Nov 2010 real.ml Finally added the theorem REAL_LE_POW_2 = |- !x. &0 <= x pow 2. Even though REAL_LE_SQUARE is close, it's nice not to have to make trivial rewrites back and forth just to use it. Thu 28th Oct 2010 Library/rstc.ml Noticed that the theorems RTC_INDUCT_L and RTC_INDUCT_R both had the hypothesis (!x y. R x y ==> P x y), which is redundant since it follows from the other two. So I just removed it; as far as I can see there are no proof changes needed. Thu 28th Oct 2010 arith.ml, passim Joe Hurd pointed out that EQ_SUC proved in arith.ml was just a duplicate of the theorem SUC_INJ from nums.ml, so I removed its proof and replaced EQ_SUC by SUC_INJ everywhere it was used. Fri 8th Oct 2010 README Updated the README file with a few modernizations, in particular a more careful discussion of checkpointing options in the light of the fact that it seems hard to get ckpt working on recent Linuxes. Fri 8th Oct 2010 theorems.ml Joe Hurd pointed out that EQ_REFL_T is a duplicate of REFL_CLAUSE, so I removed EQ_REFL_T. Fri 8th Oct 2010 Makefile, hol.ml There was a problem pointed out by Anthony V. Pulido with OCaml 3.12 because a couple of tests had equality tests for 3.10 or 3.11. Replaced this with a simple test of the first digit after the decimal point of the OCaml version in the Makefile and a string inequality comparison in "hol.ml". Wed 9th Jun 2010 nums.ml, arith.ml Made some changes suggested by Joe Hurd to allow cleaner separation into theories by avoiding the use of addition when defining the numeral constants BIT0 and BIT1. Now those are defined directly using primitive recursion (BIT0_DEF and BIT1_DEF) and the former definitions BIT0 and BIT1 are derived from those later. Wed 24th Mar 2010 nums.ml, 100/four_squares.ml, 100/ramsey.ml, Complex/complexnumbers.ml, Examples/mccarthy.ml, Examples/multiwf.ml, Mizarlight/duality_holby.ml, Model/modelset.ml, Multivariate/clifford.ml, Multivariate/topology.ml, Rqe/asym.ml, Rqe/examples.ml, Rqe/rol.ml Made changes essentially corresponding to the steps of hol-online's "hol-preparse-patch", fixing some bugs, removing redundant material and regularizing some syntax: * Replaced a = by <=> in "100/four_squares.ml" (I'd left this out of my test suite, which is why it hadn't been spotted before). * Added a couple of missing real numeral coercions "&" in "Rqe/examples.ml". * Removed duplicate definition for degree in "Rqe/asym.ml" (it's already defined in "Library/poly.ml"). * Removed terms with antiquotation, an unused relic from HOL88 in "100/ramsey.ml". * Avoided programmatically created terms in "Examples/mccarthy.ml", "Mizarlight/duality_holby.ml" and "Rqe/rol.ml". * Added IND_SUC_SPEC to regularize new_specification syntax in "nums.ml" * Stylistic change to type definition in "Complex/complexnumbers.ml", "Examples/multiwf.ml", "Mizarlight/duality_holby.ml", "Model/modelset.ml", "Multivariate/clifford.ml", "Multivariate/topology.ml". Tue 23rd Mar 2010 lists.ml Added several miscellaneous lemmas about lists: LENGTH_TL = |- !l. ~(l = []) ==> LENGTH (TL l) = LENGTH l - 1 EL_APPEND = |- !k l m. EL k (APPEND l m) = (if k < LENGTH l then EL k l else EL (k - LENGTH l) m) EL_TL = |- !n. EL n (TL l) = EL (n + 1) l EL_CONS = |- !n h t. EL n (CONS h t) = (if n = 0 then h else EL (n - 1) t) LAST_EL = |- !l. ~(l = []) ==> LAST l = EL (LENGTH l - 1) l HD_APPEND = |- !l m. HD (APPEND l m) = (if l = [] then HD m else HD l) Tue 23rd Mar 2010 sets.ml Added SET_OF_LIST_EQ_EMPTY = |- !l. set_of_list l = {} <=> l = [] Mon 22nd Mar 2010 lists.ml Added one more simple list lemma: LAST_APPEND = |- !p q. LAST(APPEND p q) = (if q = [] then LAST p else LAST q) Sat 20th Mar 2010 lists.ml, Rqe/work_thms.ml Moved the definition of BUTLAST from Rqe into the core, and added one trivial theorem about it: BUTLAST = |- BUTLAST [] = [] /\ BUTLAST (CONS h t) = (if t = [] then [] else CONS h (BUTLAST t)) APPEND_BUTLAST_LAST = |- !l. ~(l = []) ==> APPEND (BUTLAST l) [LAST l] = l Fri 12th Mar 2010 sets.ml Added CHOOSE_SUBSET_STRONG = |- !n s. (FINITE s ==> n <= CARD s) ==> (?t. t SUBSET s /\ t HAS_SIZE n) Thu 11th Mar 2010 ints.ml Made a further small stylistic change following hol-online to make the integer type bijection at the top level instead of hiding it. This introduces a new theorem int_tybij. Tue 9th Mar 2010 lists.ml Added a couple of trivial lemmas about MAP: MAP_ID = |- !l. MAP (\x. x) l = l MAP_I = |- MAP I = I Mon 8th Mar 2010 bool.ml, ind_types.ml, cart.ml, int.ml, realax.ml Made minor stylistic changes following hol-online. These essentially make it easier to recognize parsing directives and definitions. Mon 8th Mar 2010 Proofrecording/diffs/basics.ml [new file] Chantal Keller pointed out that the "needs" directive is wrong in the proof-recording version where separate type, term and theorem files are used in place of fusion. Fixed this by putting the version of basics.ml without the "needs" line in the diffs directory. Sun 7th Mar 2010 passim Moved more hol-online changes upstream, adding explicit "needs" directives to the core files in the toplevel directory. Also modified the "make_database_assignment" function in "update_database.ml" so that it adds such a directive at the top of any file it creates. Sat 6th Mar 2010 Examples/vitali.ml [new file] Added a simple example, the existence of a non-measurable set. Since it doesn't seem likely to be useful for anything, I didn't put it in the Multivariate measure theory itself. Tue 2nd Mar 2010 class.ml Jeremy Bem had pointed out that TAUT has an implicit dependency on the current default rewrites because of the embedded REWRITE_TAC[]. This can make it do more than it should if the rewrites are expanded, or stop working if they are contracted. Fixed this by partially evaluating the REWRITE_TAC[] in the definition, though this means defining it twice in class.ml, once as a sort of bootstrapping version without the COND rewrites, and one after. Mon 1st Mar 2010 update_database.ml Added alphabetic sorting of the results to the "search" function. Note that the reference variable "theorems" itself is not sorted, but sorting is imposed afterwards on the filtered results (likewise in "make_database_assignment"). Thu 25th Feb 2010 real.ml, int.ml Added the definitions of the real and integer signum function. I was tempted to add the interface "sgn" for both of them, but perhaps one doesn't want to be deprived of that as a variable name. New theorems (and of course INT_OF_REAL_THM is updated): real_sgn REAL_ABS_SGN REAL_SGN REAL_SGN_0 REAL_SGN_ABS REAL_SGN_DIV REAL_SGN_INV REAL_SGN_MUL REAL_SGN_NEG int_sgn int_sgn_th INT_ABS_SGN INT_SGN INT_SGN_0 INT_SGN_ABS INT_SGN_MUL INT_SGN_NEG Thu 25th Feb 2010 Examples/sos.ml Replaced a couple of explicit uses of the Empty constructor for finite partial functions by "undefined" and "is_undefined". This was causing a clash with the Empty constructor for the tbl type in the new update_database.ml, and avoiding the constructors is good practice anyway (maybe I should make it an abstract type). Wed 24th Feb 2010 basics.ml Tweaked "subst" to filter out trivial substitutions from the instantiation list first. This is a small thing but was actually ultimately responsible for some extreme slowness of UNWIND_CONV on very large terms owing to the underlying REWR_CONV SWAP_EXISTS_THM. For example defining the following let test n = let t = `m < n + 1` in let tm = funpow n (fun x -> mk_conj(x,t)) t in let et = list_mk_exists([`m:num`;`n:num`],tm) in time (REWR_CONV SWAP_EXISTS_THM) et; 1;; then "test 100000" used to take 9.5 seconds, and now takes 0.5. There are still some wasteful things going, so it might be worth explicitly optimizing UNWIND_CONV, but this is much better than before and might have beneficial effects elsewhere too. Fri 19th Feb 2010 update_database.ml, Examples/update_database.ml [new file] Moved the old update_database.ml script back into Examples and installed instead the somewhat nicer one from Roland Zumkeller, which doesn't rely on the OCaml sources. Fri 19th Feb 2010 sets.ml Added some theorems showing that surjectivity is exactly equivalent to inserting the function in universal and existential quantifiers and gives rise to a related property for set comprehensions: SURJECTIVE_FORALL_THM = |- !f. (!y. ?x. f x = y) <=> (!P. (!x. P (f x)) <=> (!y. P y)) SURJECTIVE_EXISTS_THM = |- !f. (!y. ?x. f x = y) <=> (!P. (?x. P (f x)) <=> (?y. P y)) SURJECTIVE_IMAGE_THM = |- !f. (!y. ?x. f x = y) <=> (!P. IMAGE f {x | P (f x)} = {x | P x}) Fri 19th Feb 2010 lists.ml Added one trivial lemma and a couple of theorems showing that injectivity and surjectivity are preserved by the MAP construct: MAP_EQ_NIL = |- !f l. MAP f l = [] <=> l = [] INJECTIVE_MAP = |- !f. (!l m. MAP f l = MAP f m ==> l = m) <=> (!x y. f x = f y ==> x = y) SURJECTIVE_MAP = |- !f. (!m. ?l. MAP f l = m) <=> (!y. ?x. f x = y) Tue 16th Feb 2010 Proofrecording/README Proofrecording/diffs/equal.ml, Proofrecording/diffs/hol.ml, Proofrecording/diffs/proofobjects_dummy.ml, Proofrecording/diffs/proofobjects_init.ml, Proofrecording/diffs/proofobjects_trt.ml, Proofrecording/hol_light/Makefile, Proofrecording/diffs/depgraph.ml [new file], Proofrecording/diffs/proofobjects_coq.ml [new file] Installed several updates from Chantal Keller to the proof-recording infrastructure, slightly tweaked and merged with my latest updates. This fixes a few incompatibilities and also extends the proof-recording infrastructure to produce Coq proofs (set HOLPROOFOBJECTS=COQ). The new code uses the ocamlgraph library, so there is a new entry "top" in the Makefile to make it. ********************** FIRST GOOGLE CODE VERSION ********************** Fri 12th Feb 2010 passim Moved many of the files from Examples into a new Library directory, and changed the names of various others, to support easier native-code compilation. Most of these changes are again modelled on Jeremy Bem's hol1process.sh file, though a few of the new names are different. The basic idea is to make the names of files distinct from each other and give rise to valid module names that are distinct from built-in OCaml modules, while moreover avoiding directory dependencies (e.g. between files in Examples and Multivariate). Here's a complete list of the changes: ind-defs.ml -> ind_defs.ml ind-types.ml -> ind_types.ml iter.ml -> iterate.ml list.ml -> lists.ml num.ml -> nums.ml sys.ml -> system.ml 100/2squares.ml -> 100/two_squares.ml 100/4squares.ml -> 100/four_squares.ml 100/agm.ml -> 100/arithmetic_geometric_mean.ml Complex/complex.ml -> Complex/complexnumbers.ml Complex/grobner.ml -> Complex/complex_grobner.ml Complex/real.ml -> Complex/complex_real.ml Complex/transc.ml -> Complex/complex_transc.ml Examples/agm.ml -> Library/agm.ml Examples/analysis.ml -> Library/analysis.ml Examples/binary.ml -> Library/binary.ml Examples/binomial.ml -> Library/binomial.ml Examples/calc_real.ml -> Library/calc_real.ml Examples/card.ml -> Library/card.ml Examples/floor.ml -> Library/floor.ml Examples/integer.ml -> Library/integer.ml Examples/isum.ml -> Library/isum.ml Examples/iter.ml -> Library/iter.ml Examples/lagrange.ml -> Examples/lagrange_lemma.ml Examples/multiplicative.ml -> Library/multiplicative.ml Examples/permutations.ml -> Library/permutations.ml Examples/pocklington.ml -> Library/pocklington.ml Examples/poly.ml -> Library/poly.ml Examples/pratt.ml -> Library/pratt.ml Examples/prime.ml -> Library/prime.ml Examples/primitive.ml -> Library/primitive.ml Examples/products.ml -> Library/products.ml Examples/rstc.ml -> Library/rstc.ml Examples/transc.ml -> Library/transc.ml Examples/update_database.ml -> update_database.ml Examples/wo.ml -> Library/wo.ml Jordan/num_calc_simp.ml -> Rqe/num_calc_simp.ml Minisat/SatSolvers.ml -> Minisat/sat_solvers.ml Minisat/dimacsTools.ml -> Minisat/dimacs_tools.ml Minisat/minisatParse.ml -> Minisat/minisat_parse.ml Minisat/minisatProve.ml -> Minisat/minisat_prove.ml Minisat/minisatResolve.ml -> Minisat/minisat_resolve.ml Minisat/satCommonTools.ml -> Minisat/sat_common_tools.ml Minisat/satScript.ml -> Minisat/sat_script.ml Minisat/satTools.ml -> Minisat/sat_tools.ml Model/set.ml -> Model/modelset.ml Multivariate/analysis.ml -> Multivariate/derivatives.ml Multivariate/complex.ml -> Multivariate/complexes.ml Multivariate/real.ml -> Multivariate/realanalysis.ml Multivariate/transc.ml -> Multivariate/transcendentals.ml Rqe/lib.ml -> Rqe/rqe_lib.ml Rqe/list.ml -> Rqe/rqe_list.ml Rqe/main.ml -> Rqe/rqe_main.ml Rqe/num.ml -> Rqe/rqe_num.ml Rqe/real.ml -> Rqe/rqe_real.ml Rqe/tactics-ext.ml -> Rqe/rqe_tactics_ext.ml Fri 12th Feb 2010 Examples/sos.ml, Minisat/minisatParse.ml, Rqe/main.ml, Rqe/matinsert.ml, Unity/mk_unless.ml Made some small changes based on the hol1process.sh file in Jeremy Bem's hol-online, mainly connected with eliminating unused state or dead code. Also removed the following files that are not runnable as they stand and were probably not intended to be in the final version: Complex/maple.ml Complex/geom.ml Boyer_Moore/induct_then.ml Boyer_Moore/testset/eval.ml Mon 25th Jan 2010 sets.ml Added the following, which is really just plugging two existing results together, but in a way that's not completely trivial for the built-in automation: HAS_SIZE_IMAGE_INJ_EQ = |- !f s n. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> (IMAGE f s HAS_SIZE n <=> s HAS_SIZE n) Tue 12th Jan 2010 sets.ml Added some trivial but useful theorems about infiniteness of various real intervals. FINITE_REAL_INTERVAL = |- (!a. ~FINITE {x | a < x}) /\ (!a. ~FINITE {x | a <= x}) /\ (!b. ~FINITE {x | x < b}) /\ (!b. ~FINITE {x | x <= b}) /\ (!a b. FINITE {x | a < x /\ x < b} <=> b <= a) /\ (!a b. FINITE {x | a <= x /\ x < b} <=> b <= a) /\ (!a b. FINITE {x | a < x /\ x <= b} <=> b <= a) /\ (!a b. FINITE {x | a <= x /\ x <= b} <=> b <= a) as well as the trivial consequence real_INFINITE : thm = |- INFINITE(:real) Fri 8th Jan 2010 pair.ml When rerunning some files I noticed a subtle issue in the benignity checking of definitions: when recreating the desired definition from the old one, it is only quantified over the explicitly quantified variables in the input term. Changed this to finally generalize afterwards over all the (remaining) free variables, which makes it completely identical with the theorem returned by the original invocation of the same definition. Fri 8th Jan 2010 sets.ml, passim Modified SET_TAC so that it throws away the assumption list. Until recently I had been under the impression that this was the actual behaviour, and was using ASM SET_TAC explicitly when I wanted the assumptions included. Needless to say, there were a few places where I'd relied on the old behaviour, so I changed SET_TAC to ASM SET_TAC in those places. Fri 8th Jan 2010 parser.ml Added an additional clause to the toplevel function "parse_preterm" so that it accepts an identifier (i.e. anything except a reserved word) as the entire string to be parsed. This was mainly intended as a convenience for the search function, so that one can do, say search [`UNION`];; instead of the slightly longer-to-type search [`(UNION)`];; Mon 28th Dec 2009 sets.ml Added a few theorems about injectivity and surjectivity of the image under a map: INJECTIVE_ON_IMAGE = |- !f u. (!s t. s SUBSET u /\ t SUBSET u /\ IMAGE f s = IMAGE f t ==> s = t) <=> (!x y. x IN u /\ y IN u /\ f x = f y ==> x = y) INJECTIVE_IMAGE = |- !f. (!s t. IMAGE f s = IMAGE f t ==> s = t) <=> (!x y. f x = f y ==> x = y) SURJECTIVE_ON_IMAGE = |- !f u v. (!t. t SUBSET v ==> (?s. s SUBSET u /\ IMAGE f s = t)) <=> (!y. y IN v ==> (?x. x IN u /\ f x = y)) SURJECTIVE_IMAGE = |- !f. (!t. ?s. IMAGE f s = t) <=> (!y. ?x. f x = y) Sun 27th Dec 2009 sets.ml Added FINITE_HAS_SIZE = |- !s. FINITE s <=> s HAS_SIZE CARD s and generalized the theorem FINITE_PRODUCT_DEPENDENT from pairs to the application of any function: FINITE_PRODUCT_DEPENDENT = |- !f s t. FINITE s /\ (!x. x IN s ==> FINITE (t x)) ==> FINITE {f x y | x IN s /\ y IN t x} Fri 20th Nov 2009 sets.ml Noticed that FINITE_UNIONS could be stronger (funny that I only just noticed this). Changed the name of the former FINITE_UNIONS to FINITE_FINITE_UNIONS and made FINITE_UNIONS the stronger theorem. I hope that most proofs will not break after this change because it's usually used with SIMP_TAC. FINITE_FINITE_UNIONS = |- !s. FINITE s ==> (FINITE (UNIONS s) <=> (!t. t IN s ==> FINITE t)) FINITE_UNIONS = |- !s. FINITE (UNIONS s) <=> FINITE s /\ (!t. t IN s ==> FINITE t) Fri 20th Nov 2009 Examples/floor.ml Added RATIONAL_ALT for use in some places where I'd used that as an adhoc definition (100/constructible.ml and 100/liouville.ml). RATIONAL_ALT = |- !x. rational x <=> (?p q. ~(q = 0) /\ abs x = &p / &q) Tue 17th Nov 2009 Examples/floor.ml Added a definition of "rational" and closure theorems, corresponding as appropriate to those for "integer". Mon 16th Nov 2009 real.ml Added a suite of theorems to complement REAL_POW_LE2_ODD (which is still there but with a simpler proof, the complicated one now reserved for REAL_POW_LT2_ODD). REAL_POW_LT2_ODD = |- !n x y. x < y /\ ODD n ==> x pow n < y pow n REAL_POW_LT2_ODD_EQ = |- !n x y. ODD n ==> (x pow n < y pow n <=> x < y) REAL_POW_LE2_ODD_EQ = |- !n x y. ODD n ==> (x pow n <= y pow n <=> x <= y) REAL_POW_EQ_ODD_EQ = |- !n x y. ODD n ==> (x pow n = y pow n <=> x = y) REAL_POW_EQ_ODD = |- !n x y. ODD n /\ x pow n = y pow n ==> x = y REAL_POW_EQ_EQ;; |- !n x y. x pow n = y pow n <=> (if EVEN n then n = 0 \/ abs x = abs y else x = y) Wed 11th Nov 2009 sets.ml Added PAIRWISE_MONO = |- !r s t. pairwise r s /\ t SUBSET s ==> pairwise r t Tue 3rd Nov 2009 Proofrecording/diffs/thm.ml Added a new version of thm.ml from Chantal Keller, which makes things basically compatible with the new core, though it's still implemented via separate type.ml, term.ml and thm.ml files instead of a single fusion.ml; at some point I should go through and make it all perfectly compatible. Wed 21st Oct 2009 Examples/primitive.ml [new file] Added a file giving the usual existence results for primitive roots. Mon 19th Oct 2009 theorems.ml Added universally quantified versions of unwinding theorems, since I often seem to use these and have to laboriously use SIMP_TAC[] then LEFT_FORALL_IMP_THM and EXISTS_REFL. FORALL_UNWIND_THM1 = |- !P a. (!x. a = x ==> P x) <=> P a FORALL_UNWIND_THM2 = |- !P a. (!x. x = a ==> P x) <=> P a Wed 14th Oct 2009 Examples/integer.ml Added a couple of theorems to help transfer solutions of congruences back to the natural numbers: INT_LINEAR_CONG_POS = |- !n a x. ~(n = &0) ==> (?y. &0 <= y /\ (a * x == a * y) (mod n)) INT_CONG_SOLVE_POS = |- !a b n. coprime (a,n) /\ ~(n = &0 /\ abs a = &1) ==> (?x. &0 <= x /\ (a * x == b) (mod n)) Tue 13th Oct 2009 int.ml Added outer quantifiers to these two theorems: INT_FORALL_POS = |- !P. (!n. P(&n)) <=> (!i. &0 <= i ==> P i) INT_EXISTS_POS = |- !P. (?n. P(&n)) <=> (?i. &0 <= i /\ P i) and added the following two analogous ones: INT_FORALL_ABS = |- !P. (!n. P(&n)) <=> (!x. P (abs x)) INT_EXISTS_ABS = |- !P. (?n. P(&n)) <=> (?x. P (abs x)) Tue 13th Oct 2009 Examples/pratt.ml, Examples/pocklington.ml Renamed PHI_PRIME to PHI_PRIME_EQ PHI_PRIME_EQ = |- !n. phi n = n - 1 /\ ~(n = 0) /\ ~(n = 1) <=> prime n and made PHI_PRIME the following: PHI_PRIME = |- !p. prime p ==> phi p = p - 1 Tue 13th Oct 2009 sets.ml Added the following, which is a bit more convenient than manually chaining two lemmas: CARD_SUBSET_IMAGE = |- !f s t. FINITE t /\ s SUBSET IMAGE f t ==> CARD s <= CARD t Tue 13th Oct 2009 Examples/update_database.ml Propagated an earlier fix from the "search" implementation in "help.ml", to ignore pure variables in patterns. Mon 12th Oct 2009 iter.ml Added the following about finiteness of integer segments; although trivial-looking it's actually quite a bit of work to derive from scratch: FINITE_INTSEG = |- (!l r. FINITE {x:int | l <= x /\ x <= r}) /\ (!l r. FINITE {x:int | l <= x /\ x < r}) /\ (!l r. FINITE {x:int | l < x /\ x <= r}) /\ (!l r. FINITE {x:int | l < x /\ x < r}) Thu 27th Aug 2009 sets.ml Added two more lemmas about subsets of images: EXISTS_SUBSET_IMAGE = |- !P f s. (?t. t SUBSET IMAGE f s /\ P t) <=> (?t. t SUBSET s /\ P (IMAGE f t)) EXISTS_FINITE_SUBSET_IMAGE = |- !P f s. (?t. FINITE t /\ t SUBSET IMAGE f s /\ P t) <=> (?t. FINITE t /\ t SUBSET s /\ P (IMAGE f t)) Sun 23rd Aug 2009 sets.ml Added CROSS_EQ_EMPTY = |- !s t. s CROSS t = {} <=> s = {} \/ t = {} Sat 22nd Aug 2009 iter.ml Added two general number-segment theorems for the general "iterate": ITERATE_CLAUSES_NUMSEG = |- !op. monoidal op ==> (!m. iterate op (m..0) f = (if m = 0 then f 0 else neutral op)) /\ (!m n. iterate op (m..SUC n) f = (if m <= SUC n then op (iterate op (m..n) f) (f (SUC n)) else iterate op (m..n) f)) ITERATE_PAIR = |- !op. monoidal op ==> (!f m n. iterate op (2 * m..2 * n + 1) f = iterate op (m..n) (\i. op (f (2 * i)) (f (2 * i + 1)))) and two instances of the latter: NSUM_PAIR = |- !f m n. nsum (2 * m..2 * n + 1) f = nsum (m..n) (\i. f (2 * i) + f (2 * i + 1)) SUM_PAIR = |- !f m n. sum (2 * m..2 * n + 1) f = sum (m..n) (\i. f (2 * i) + f (2 * i + 1)) Fri 21st Aug 2009 sets.ml Added the following for dealing with unions and intersections of images (explicit or implicit), which can otherwise be a bit tedious: UNIONS_IMAGE = |- !f s. UNIONS (IMAGE f s) = {y | ?x. x IN s /\ y IN f x} INTERS_IMAGE = |- !f s. INTERS (IMAGE f s) = {y | !x. x IN s ==> y IN f x} UNIONS_GSPEC = |- (!P f. UNIONS {f x | P x} = {a | ?x. P x /\ a IN f x}) /\ (!P f. UNIONS {f x y | P x y} = {a | ?x y. P x y /\ a IN f x y}) /\ (!P f. UNIONS {f x y z | P x y z} = {a | ?x y z. P x y z /\ a IN f x y z}) INTERS_GSPEC = |- (!P f. INTERS {f x | P x} = {a | !x. P x ==> a IN f x}) /\ (!P f. INTERS {f x y | P x y} = {a | !x y. P x y ==> a IN f x y}) /\ (!P f. INTERS {f x y z | P x y z} = {a | !x y z. P x y z ==> a IN f x y z}) Tue 18th Aug 2009 sets.ml Removed the theorem FINITE_SUBSETS, which is just a duplicate of FINITE_POWERSET. Mon 17th Aug 2009 sets.ml Added the following, which is somewhat trivial but quite laborious to derive: SET_PROVE_CASES = |- !P. P {} /\ (!a s. ~(a IN s) ==> P (a INSERT s)) ==> (!s. P s) Mon 10th Aug 2009 basics.ml Norbert Voelker pointed out that list_mk_icomb can give unexpected "capture" of type variables because of the iterative implementation instantiating part at a time using mk_icomb. For example list_mk_icomb "," [`b:B`;`c:C`];; formerly gave `(b:C,c:C)`. I changed to a different implementation that computes the type instantiation once and for all. Sat 8th Aug 2009 iter.ml Got rid of ITERATE_CLOSED_GEN and made simple ITERATE_CLOSED still stronger without any finiteness hypothesis: ITERATE_CLOSED = |- !op. monoidal op ==> (!P. P(neutral op) /\ (!x y. P x /\ P y ==> P(op x y)) ==> (!f s. (!x. x IN s /\ ~(f x = neutral op) ==> P(f x)) ==> P(iterate op s f))) Also added two special cases: NSUM_CLOSED = |- !P f s. P 0 /\ (!x y. P x /\ P y ==> P(x + y)) /\ (!a. a IN s ==> P(f a)) ==> P(nsum s f) SUM_CLOSED = |- !P f s. P(&0) /\ (!x y. P x /\ P y ==> P(x + y)) /\ (!a. a IN s ==> P(f a)) ==> P(sum s f) Sun 26th Jul 2009 class.ml Shortened the proof of EXCLUDED_MIDDLE using a method shown to me by Mark Adams, which uses a smaller and simpler select-term in the main step. Thu 23rd Jul 2009 hol.ml Roland Zumkeller pointed out to me that "fusion.ml" is loaded by #use, whereas it should be loads to be consistent with the others and make it work even when not in the current directory. Tue 21st Jul 2009 iter.ml Removed the unnecessary finiteness hypotheses from NSUM_RESTRICT_SET = |- !P s f. nsum {x | x IN s /\ P x} f = nsum s (\x. if P x then f x else 0) SUM_RESTRICT_SET = |- !P s f. sum {x | x IN s /\ P x} f = sum s (\x. if P x then f x else &0) Tue 21st Jul 2009 iter.ml Added another useful lemma for computing slightly tweaked sums: SUM_CASES_1 = |- !s a. FINITE s /\ a IN s ==> sum s (\x. if x = a then y else f x) = sum s f + y - f a Mon 20th Jul 2009 iter.ml Added the following; even though it's rarely needed the derivation is surprisingly painful: ITERATE_OP_GEN = |- !op. monoidal op ==> (!f g s. FINITE (support op f s) /\ FINITE (support op g s) ==> iterate op s (\x. op (f x) (g x)) = op (iterate op s f) (iterate op s g)) and then the two usual instances: NSUM_ADD_GEN = |- !f g s. FINITE {x | x IN s /\ ~(f x = 0)} /\ FINITE {x | x IN s /\ ~(g x = 0)} ==> nsum s (\x. f x + g x) = nsum s f + nsum s g SUM_ADD_GEN = |- !f g s. FINITE {x | x IN s /\ ~(f x = &0)} /\ FINITE {x | x IN s /\ ~(g x = &0)} ==> sum s (\x. f x + g x) = sum s f + sum s g Wed 8th Jul 2009 ind-defs.ml Added derive_strong_induction, which is pretty much the same thing as the function of that name in Tom Melham's original inductive definitions package. In fact, this was about the only significant respect in which my package was not a generalization of it. This was inspired by Sean McLaughlin's observation that the Coq inductive definitions package gives you this automatically, and it can be quite useful. Tue 7th Jul 2009 class.ml Freek and Cezary pointed out that I have duplicates of the two basic_rewrites |- ~T <=> F and |- ~F <=> T, because I add NOT_CLAUSES_WEAK and then later NOT_CLAUSES. Changed the latter to CONJUNCT1 NOT_CLAUSES. Mon 6th Jul 2009 pair.ml Added the following, for constructs that are otherwise painful: FORALL_PAIRED_THM = |- !P. (!(x,y). P x y) <=> (!x y. P x y) EXISTS_PAIRED_THM = |- !P. (?(x,y). P x y) <=> (?x y. P x y) FORALL_TRIPLED_THM = |- !P. (!(x,y,z). P x y z) <=> (!x y z. P x y z) EXISTS_TRIPLED_THM = |- !P. (?(x,y,z). P x y z) <=> (?x y z. P x y z) Also added outer quantifiers to FORALL_PAIR_THM, EXISTS_PAIR_THM and LAMBDA_PAIR_THM, also changing the variable from P to t in the last. Tue 30th Jun 2009 tactics.ml Dan Synek pointed out some issues over using REPEAT_GTCL on IMP_RES_THEN. I decided that things like this would work properly if IMP_RES_THEN checked its list of derived tactics to make sure it's nonempty, so that it will actually fail in such cases instead of running an empty tactic. So I made this change in ANTE_RES_THEN and IMP_RES_THEN. Tue 30th Jun 2009 hol.ml Sean McLaughlin pointed out that relocation of hol_dir also spoils the list of loaded files in "loaded_files", which uses absolute paths. I changed this so that the filenames just have the "basename", and it seems to work well. Sat 27th Jun 2009 make.ml Changed "checkpoint" to use the "-l" option to CryoPID, since it does seem to give at least some measure of portability in the resulting binaries, though problems over exceptions and general unpredictability remain. Sat 27th Jun 2009 hol.ml, help.ml Sean McLaughlin pointed out that the "help" function doesn't adjust dynamically to the change of !hol_dir, so I changed it to redo more stuff dynamically per argument. It also seemed worth making help_path an assignable reference variable so people can add more help entries (maybe indeed I should do this for some non-core extensions). I was now also dissatisfied that load_path needs changing separately. So I set up a scheme where a new function "hol_expand_directory" is applied dynamically to help_path and load_path, inside "help" and "file_on_path" (hence derivatives like "loadt"). It maps an initial $ into !hol_dir. Tue 16th Jun 2009 100/lagrange.ml [new file] Freek was asking me about Lagrange's theorem, and I realized I don't have it in the standard system anywhere (this and a few others oddments were in Work/group.ml). So I put it here. Tue 16th Jun 2009 help.ml Sean McLaughlin pointed out that the use of "cat" in the help system won't work on Windows unless you are inside Cygwin. It wasn't really needed anyway to do "cat file | sed -f script": I replaced it by "sed -f script file". Mon 1st Jun 2009 Examples/multiwf.ml Renamed two theorems here that now overwrite other theorems in the main system: FINITE_EMPTY->EMPTY_IS_FINITE and FINITE_SING->SING_IS_FINITE. Tue 26th May 2009 Ntrie/ntrie.ml, Ntrie/ntrie_tests.ml [new files] Added a nice new library by Clelia Lomuto and Marco Maggesi that performs efficient set operations on sets of natural numbers using a trie representation internally. Mon 25th May 2009 Complex/cpoly.ml, Complex/fundamental.ml Added more "needs" directives to these files so that you can load more things individually without special measures. This was inspired by a bug report from Loic Pottier who tried to load Complex/grobner.ml directly. Mon 25th May 2009 pair.ml Added an outer universal quantifier over P to the theorem pair_INDUCT; Norbert Voelker had pointed out that it didn't fit the pattern of all the other induction theorems that have this quantifier. Sat 16th May 2009 help.ml Modified "search" function so it ignores pure variables, with a warning, and also returns the empty list if the input list is nonempty but entirely full of variables. Sat 16th May 2009 real.ml, int.ml Added the following two to complete the suite with LE and LT variants: REAL_EQ_SQUARE_ABS = |- !x y. abs x = abs y <=> x pow 2 = y pow 2 INT_EQ_SQUARE_ABS = |- !x y. abs x = abs y <=> x pow 2 = y pow 2 Thu 14th May 2009 sets.ml Added the following, which is potentially quite useful for otherwise tedious reasoning patterns: SURJECTIVE_IMAGE_EQ = |- !s t. (!y. y IN t ==> (?x. f x = y)) /\ (!x. f x IN t <=> x IN s) ==> IMAGE f s = t Sat 9th May 2009 Examples/update_database.ml Made a small tweak so that "make_database_assignment" applies "uniq" to its list. Some duplicates do otherwise happen, and now they are mostly deliberate overwritings like num_Axiom. Tue 5th May 2009 cart.ml Added a couple of theorems that are trivial, but can be used to eliminate conditions on indices in some interesting situations, e.g. in vector arithmetic operations. CART_EQ_FULL = |- !x y. x = y <=> (!i. x$i = y$i) FINITE_INDEX_INRANGE = |- !i. ?k. 1 <= k /\ k <= dimindex(:N) /\ (!x:A^N. x$i = x$k) Sat 2nd May 2009 sets.ml Added a trivial theorem that I sometimes want: FINITE_EMPTY = |- FINITE {} Of course, it already exists, but CONJUNCT1 FINITE_RULES is slightly more to type, while just FINITE_RULES can make the simplifier work much harder trying to backchain. Mon 20th Apr 2009 sets.ml Added a couple of potentially useful theorems; there are similar things in Examples/card.ml, but not the "universal" form and it seems natural for it to be in the core. BIJECTIVE_ON_LEFT_RIGHT_INVERSE = |- !f s t. (!x. x IN s ==> f x IN t) ==> ((!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) /\ (!y. y IN t ==> (?x. x IN s /\ f x = y)) <=> (?g. (!y. y IN t ==> g y IN s) /\ (!y. y IN t ==> f (g y) = y) /\ (!x. x IN s ==> g (f x) = x))) BIJECTIVE_LEFT_RIGHT_INVERSE = |- !f. (!x y. f x = f y ==> x = y) /\ (!y. ?x. f x = y) <=> (?g. (!y. f (g y) = y) /\ (!x. g (f x) = x)) Fri 17th Apr 2009 sets.ml Added two degenerate cases for "pairwise", which are trivial to prove but where you wouldn't really want to divert a proof to do so: PAIRWISE_EMPTY = |- !r. pairwise r {} <=> T PAIRWISE_SING = |- !r x. pairwise r {x} <=> T Tue 14th Apr 2009 equal.ml Changed the implementation of CACHE_CONV slightly so that it pushes the function ALPHA_HACK inside the term net. The motivation was that Freek discovered a problem when adding definitions to constants in theorems: in the existing implementation this was comparing theorems when inserting things into the list at a particular net node, when used in the real arithmetic decision procedure. Actually, since it's almost always used with functions now, maybe I should just simplify the code by forgetting about eliminating duplicates in this list. Sat 11th Apr 2009 pa_j_3.11.ml (new file) Just to avoid generating a harmless error message, copied "pa_j_3.10.ml" to "pa_j_3.11.ml". Fortunately no changes are needed for the latest camlp5 (I just tried camlp5 5.11 and it's fine) so maybe thanks to camlp5 I'll be able to stop continually having to release new syntax extension files for each major OCaml release. Tue 31st Mar 2009 Examples/moebius.ml Added a new file Examples/moebius.ml by Gianni Ciolli, Graziano Gentili and Marco Maggesi defining Moebius functions and proving they are the biholomorphisms of the unit disc. This uses Cartan's theorem, which they also contributed and I installed in the Multivariate theories. Tue 17th Mar 2009 basics.ml Added "mk_let", which Roland Zumkeller had pointed out was missing, even though there was a dead link to documentation for it (of course, I now added the documentation too). Also made dest_let a bit fussier: it used not to check that the head operator was in fact "LET". Tue 24th Feb 2009 basics.ml, ind-defs.ml Fixed a bug in subst pointed out by Mark Adams, where for example: subst [v,`1`] (mk_exists(v,`1 + 1 = 2`));; where v happens to be the next genvar to be generated. Just used "variants" to avoid the existing variables in the term when creating intermediate variables. (To do this, I moved "variables" from ind-defs.ml, which is hardly a natural place for it anyway, and moved back the definition of "variants" within "basics.ml".) There must be a few other places where such genvar coincidences would kill things too. Sun 22nd Feb 2009 sets.ml Finally added these, which I seem to generate quite a bit at the moment: IMAGE_ID = |- !s. IMAGE (\x. x) s = s IMAGE_I = |- !s. IMAGE I s = s Fri 20th Feb 2009 sets.ml Added natural duals to corresponding FORALL_xxx theorems: EXISTS_IN_INSERT = |- !P a s. (?x. x IN a INSERT s /\ P x) <=> P a \/ (?x. x IN s /\ P x) EXISTS_IN_GSPEC = |- (!P f. (?z. z IN {f x | P x} /\ Q z) <=> (?x. P x /\ Q (f x))) /\ (!P f. (?z. z IN {f x y | P x y} /\ Q z) <=> (?x y. P x y /\ Q (f x y))) /\ (!P f. (?z. z IN {f w x y | P w x y} /\ Q z) <=> (?w x y. P w x y /\ Q (f w x y))) Fri 13th Feb 2009 sets.ml Added IMAGE_INTER_INJ = |- !f s t. (!x y. f x = f y ==> x = y) ==> IMAGE f (s INTER t) = IMAGE f s INTER IMAGE f t and also added outer universal quantifiers to similar theorems IMAGE_DIFF_INJ and IMAGE_DELETE_INJ. Thu 15th Jan 2009 Examples/borsuk.ml Added a proof of the Borsuk-Ulam theorem, so far just for the ordinary 2-sphere in real^3. Wed 14th Jan 2009 sets.ml Added FUN_IN_IMAGE = |- !f s x. x IN s ==> f x IN IMAGE f s Wed 14th Jan 2009 sets.ml Added PSUBSET_ALT = |- !s t:A->bool. s PSUBSET t <=> s SUBSET t /\ (?a. a IN t /\ ~(a IN s)) Thu 25th Dec 2008 sets.ml Added SET_OF_LIST_MAP = |- !f l. set_of_list (MAP f l) = IMAGE f (set_of_list l) Wed 24th Dec 2008 list.ml Added the following useful theorem, which is actually one of very few about EL: MEM_EXISTS_EL = |- !l x. MEM x l <=> (?i. i < LENGTH l /\ x = EL i l) Sun 21st Dec 2008 sets.ml Added the following, not in wf.ml because it needs set notions like "FINITE": WF_FINITE = |- !(<<). (!x. ~(x << x)) /\ (!x y z. x << y /\ y << z ==> x << z) /\ (!x. FINITE {y | y << x}) ==> WF (<<) Sun 21st Dec 2008 arith.ml, sets.ml Moved the following theorems from "sets.ml" back to "arith.ml": TRANSITIVE_STEPWISE_LT_EQ, TRANSITIVE_STEPWISE_LT, TRANSITIVE_STEPWISE_LE_EQ, TRANSITIVE_STEPWISE_LE. This was mainly so I could exploit them sooner, but in any case they seem to belong there more naturally. Sat 20th Dec 2008 Examples/card.ml Added a few more general lemmas and particularly a few non-trivial results about countable sets. Sat 20th Dec 2008 sets.ml, Examples/card.ml It seemed ugly having these hardly-used relations like CARD_GE that are then superseded by >=_c etc. in Examples/card.ml. So I replaced the definitions in sets.ml with those <=_c forms, adding >= and >, also moving LE_C and proving GE_C which is essentially a replacement for the old CARD_GE. Deleted the little-used theorems CARD_GE_REFL, CARD_GE_TRANS and FINITE_HAS_SIZE_LEMMA (these would belong in Examples/card.ml) but put copies in Jordan for now. Fri 19th Dec 2008 sets.ml Added the following, which is actually often nicer than using SIMPLE_IMAGE, and applicable to cases with 1, 2 or 3 parameters: FORALL_IN_GSPEC = |- (!P f. (!z. z IN {f x | P x} ==> Q z) <=> (!x. P x ==> Q (f x))) /\ (!P f. (!z. z IN {f x y | P x y} ==> Q z) <=> (!x y. P x y ==> Q (f x y))) /\ (!P f. (!z. z IN {f w x y | P w x y} ==> Q z) <=> (!w x y. P w x y ==> Q (f w x y))) Tue 16th Dec 2008 Makefile Made a few small tweaks to the Makefile, in particular forcing loading Examples/update_database.ml right at the end of the loads for each image. I had earlier encountered problems with duplicates because of the explicit assignment of theorems in "multivariate_database.ml" etc. Sun 14th Dec 2008 sets.ml Added the following, which saves some trivial "unwinding" FORALL_IN_INSERT = |- !P a s. (!x. x IN a INSERT s ==> P x) <=> P a /\ (!x. x IN s ==> P x) Tue 9th Dec 2008 cart.ml Added "vector", which was formerly defined in the multivariate theories and for the special type real. Here generalized it to type A, and eliminated special choice of zero for higher elements (which is irrelevant anyway and doesn't really generalize). The immediate impetus was that I wanted to use it to define matrices. Fri 5th Dec 2008 iter.ml Added the following theorems about differences of powers. Could add the sums of powers too, but they would hardly be nicer than just instantiating one of these with a negated argument. REAL_SUB_POW = |- !x y n. 1 <= n ==> x pow n - y pow n = (x - y) * sum (0..n - 1) (\i. x pow i * y pow (n - 1 - i)) REAL_SUB_POW_R1 = |- !x n. 1 <= n ==> x pow n - &1 = (x - &1) * sum (0..n - 1) (\i. x pow i) REAL_SUB_POW_L1 = |- !x n. 1 <= n ==> &1 - x pow n = (&1 - x) * sum (0..n - 1) (\i. x pow i) Fri 5th Dec 2008 arith.ml Added EXP_ZERO = |- !n. 0 EXP n = (if n = 0 then 1 else 0) Thu 4th Dec 2008 int.ml Added another forgotten analog of a real theorem: INT_POW_ZERO = |- !n. &0 pow n = (if n = 0 then &1 else &0) Thu 4th Dec 2008 Examples/mangoldt.ml Examples/agm.ml Examples/multiplicative.ml [new files] Added three new files to Examples with stuff that is generally useful. The mangoldt.ml one is essentially common to 100/pnt.ml and 100/dirichlet.ml, so it makes sense to have it common. The agm.ml one was already basically there in 100/agm.ml, but I factored out the explicit use of roots so this version is more portable. The multiplicative.ml defines the notion for both N and R and has the Mobius function for R. Wed 4th Dec 2008 100/dirichlet.ml [new file] Added the proof of Dirichlet's theorem to the Great 100 Theorems directory. Wed 3rd Dec 2008 iter.ml Changed the quantification in NSUM_OFFSET and SUM_OFFSET. Previously one variable was left free, and it was inconvenient to have the "offset" variable at the end of the list because that is often necessary to instantiate in order to get higher order matching working. Wed 3rd Dec 2008 sets.ml Added a lemma that I seem to regenerate quite often. I'm not very keen on the name but it's consistent with EMPTY_GSPEC, which admittedly I don't use much. SING_GSPEC = |- (!a. {x | x = a} = {a}) /\ (!a. {x | a = x} = {a}) Also added the following, which I seem to want quite often: FINITE_SING = |- !a. FINITE {a} Sat 29th Nov 2008 real.ml Added a few useful theorems that are obvious holes in the current pattern: REAL_INV_POW = |- !x n. inv (x pow n) = inv x pow n REAL_INV_LT_1 = |- !x. &1 < x ==> inv x < &1 REAL_INV_1_LT = |- !x. &0 < x /\ x < &1 ==> &1 < inv x Thu 27th Nov 2008 passim Changed a lot of instances of GSYM IMP_IMP to IMP_CONJ, and likewise changed explicit TAUT calls to IMP_CONJ, IMP_CONJ_ALT. Even caught one or two instances of IMP_IMP and EQ_IMP replacing explicit tautologies. Wed 26th Nov 2008 theorems.ml Added the following, which I keep regenerating ad hoc: IMP_CONJ_ALT = |- p /\ q ==> r <=> q ==> p ==> r Wed 26th Nov 2008 iter.ml Added two theorems relating explicit number segments to the m..n notation: NUMSEG_LE = |- !n. {x | x <= n} = 0..n NUMSEG_LT = |- !n. {x | x < n} = (if n = 0 then {} else 0..n - 1) Mon 24th Nov 2008 drule.ml Changed two instances of "=" between terms (when checking whether the result needs non-trivial alpha-conversion) to Pervasives.compare, one each in PART_MATCH and GEN_PART_MATCH. This came up when checking some big tables where this step was taking half a second. Then wondered about matching itself. Although in this case it's not the bottleneck, I made some corresponding changes inside term_match (more of them). Sun 23rd Nov 2008 real.ml, Examples/isum.ml Added a version of telescoping the other way round, otherwise it's tedious to apply it. SUM_DIFFS_ALT = |- !m n. sum(m..n) (\k. f (k + 1) - f k) = (if m <= n then f (n + 1) - f m else &0) ISUM_DIFFS_ALT = |- !m n. isum(m..n) (\k. f (k + 1) - f k) = (if m <= n then f (n + 1) - f m else &0) Sun 23rd Nov 2008 int.ml Added the following, which had somehow got left out: INT_SUB_RDISTRIB = |- !x y z. (x - y) * z = x * z - y * z Sun 23rd Nov 2008 Examples/isum.ml [new file] Added a new file of theorems about integer sums. I decided to do it by generalizing INT_OF_REAL_THM, which works quite nicely, and would be quicker if I ever rolled it into the core. The generalized INT_OF_REAL_THM could conceivably be useful in itself. Sun 23rd Nov 2008 int.ml Surprisingly, found that my last change had broken INTEGER_TAC, which was sometimes calling int_ideal_cofactors in cases where not all the terms are integers. Put in filtering, but really I should think more about how this happens, since it may be indicative of some more systematic improvements I could make. Sat 22nd Nov 2008 calc_rat.ml, int.ml Modified "int_ideal_cofactors" and "real_ideal_cofactors" to give a clearer error messages if the types of the terms are unexpected. This was stimulated by a report from Loic Pottier who tried the example from the reference manual. Anyway I also added a "prioritize_real()" to the help entry for "int_ideal_cofactors". Fri 21st Nov 2008 calc_rat.ml Made a change corresponding to what I did for num's MAX and MIN: added conversions REAL_RAT_MAX_CONV and REAL_RAT_MIN_CONV and rolled them into REAL_RAT_RED_CONV and hence REAL_RAT_REDUCE_CONV. Did likewise over integers, adding conversions INT_MAX_CONV and INT_MIN_CONV and rolled them into INT_RED_CONV and hence INT_REDUCE_CONV. Also added two supporting pseudo-definitions: INT_MAX = |- !x y. max x y = (if x <= y then y else x) INT_MIN = |- !x y. min x y = (if x <= y then x else y) Fri 21st Nov 2008 Examples/schnirelmann.ml [new file] Added my little file of properties of Schnirelmann density to Examples, since it is relatively self-contained and potentially useful. Wed 19th Nov 2008 iter.ml Added the following; the "sum" analog was already there but now I needed the natural number version: NSUM_IMAGE_NONZERO = |- !d i s. FINITE s /\ (!x y. x IN s /\ y IN s /\ ~(x = y) /\ i x = i y ==> d (i x) = 0) ==> nsum (IMAGE i s) d = nsum s (d o i) Tue 18th Nov 2008 iter.ml Added a result about enumerating elements of a finite set in order; it's strange I'd never wanted this before! TOPOLOGICAL_SORT = |- !(<<). (!x y. x << y /\ y << x ==> x = y) /\ (!x y z. x << y /\ y << z ==> x << z) ==> (!n s. s HAS_SIZE n ==> (?f. s = IMAGE f (1..n) /\ (!j k. j IN 1..n /\ k IN 1..n /\ j < k ==> ~(f k << f j)))) Mon 17th Nov 2008 calc_num.ml Added NUM_MAX_CONV and NUM_MIN_CONV and rolled them into NUM_RED_CONV and NUM_REDUCE_CONV. Mon 17th Nov 2008 real.ml Added two more theorems that I was surprised weren't (quite) already there. REAL_POW_EQ_1_IMP = |- !x n. ~(n = 0) /\ x pow n = &1 ==> abs x = &1 REAL_POW_EQ_1 = |- !x n. x pow n = &1 <=> abs x = &1 /\ (x < &0 ==> EVEN n) \/ n = 0 Fri 14th Nov 2008 Makefile Changed the Makefile to load the dynamic database update into the basic hol image, not only the extra builds on top. Put a final "search" in those, however, so that the very first search is not slow. Fri 14th Nov 2008 sets.ml Added three more theorems about unions that I was rather surprised I'd never done before: CARD_UNION_GEN = |- !s t. FINITE s /\ FINITE t ==> CARD(s UNION t) = (CARD s + CARD t) - CARD(s INTER t) CARD_UNION_OVERLAP_EQ = |- !s t. FINITE s /\ FINITE t ==> (CARD(s UNION t) = CARD s + CARD t <=> s INTER t = {}) CARD_UNION_OVERLAP = |- !s t. FINITE s /\ FINITE t /\ CARD(s UNION t) < CARD s + CARD t ==> ~(s INTER t = {}) Wed 12th Nov 2008 int.ml Added a dual to the existing theorem for the universal quantifier: INT_EXISTS_POS = |- (?n. P (&n)) <=> (?i. &0 <= i /\ P i) Wed 12th Nov 2008 int.ml Changed the definition of (a == b) to assume both sides have the same type. This is a useful type constraint and I can't imagine the accidental generality ever actually being useful. Tue 11th Nov 2008 arith.ml, real.ml, int.ml, passim Finally added natural number MAX and MIN: MAX = |- !m n. MAX m n = (if m <= n then n else m) MIN = |- !m n. MIN m n = (if m <= n then m else n) and some basic supporting theorems, just enough to roll it into ARITH_RULE etc. REAL_OF_NUM_MAX = |- !m n. max (&m) (&n) = &(MAX m n) REAL_OF_NUM_MIN = |- !m n. min (&m) (&n) = &(MIN m n) INT_OF_NUM_MAX = |- !m n. max (&m) (&n) = &(MAX m n) INT_OF_NUM_MIN = |- !m n. min (&m) (&n) = &(MIN m n) Removed corresponding or slightly different definitions from the places where it was already used, and occasionally simplified proofs just to use the enhanced ARITH_TAC. Sun 2nd Nov 2008 Mizarlight/make.ml There was still an issue with getting Pa_j when not in the HOL directory, so added the line: Topdirs.dir_directory (!hol_dir);; Sun 2nd Nov 2008 Mizarlight/make.ml Changed "#load" to my own explicit path version so it works wherever HOL Light is installed. Sat 1st Nov 2008 Minisat/make.ml, Boyer_Moore/make.ml Changed "#use" into "loads" so that it will find the path wherever HOL Light is installed. Sat 1st Nov 2008 Examples/pratt.ml, Examples/pocklington.ml Added "2>/dev/null" to the GP invocations so we avoid the "reading .gprc" messages. Sat 25th Oct 2008 Makefile, Multivariate/make.ml, Multivariate/make_complex.ml, Multivariate/multivariate_database.ml [new file], Multivariate/complex_database.ml [new file] Added pre-built databases for the multivariate and complex-multivariate theories, as a convenience for people who aren't able to use the dynamic updating. Sat 25th Oct 2008 Examples/sos.ml Put preprocessing of DECIMAL (i.e. the #x.y notation) into REAL_SOS in exactly the same way as REAL_ARITH. This was motivated by the example of using REAL_SQRTSOSFIELD on `&2 * #1.255 < sqrt(&8)` that I mentioned on the Flyspeck list. Sat 25th Oct 2008 set.ml Added these two lemmas, which came up naturally in a query from Truong Nguyen on the Flyspeck list. CARD_SET_OF_LIST_LE = |- !l. CARD(set_of_list l) <= LENGTH l HAS_SIZE_SET_OF_LIST = |- !l. (set_of_list l) HAS_SIZE (LENGTH l) <=> PAIRWISE (\x y. ~(x = y)) l Fri 24th Oct 2008 drule.ml, simp.ml, tactics.ml Since the change to PART_MATCH on 22nd Sep was breaking too much stuff, I restored the old version but added a new one GEN_PART_MATCH with the new behaviour, and used that so far just in congruence rule handling in the simplifier. Tue 21st Oct 2008 Examples/pratt.ml, Examples/pocklington.ml When trying out HOL Light on a new platform (my own little Debian thing inside Virtualbox) I hit a problem with the PARI/GP interface because I was looking for the first "[" to find the output, yet in this case that occurred in the startup banner. Made the chopping of the output slightly less unintelligent by starting at the rightmost "]" then the rightmost "[". Wed 24th Sep 2008 Examples/analysis.ml Made the overload skeleton and overloading instances of "sum" more general; before the overload skeleton was restricted to sums out of a set of numbers. This was pointed out by Tom Hales. Tue 23rd Sep 2008 tactics.ml Fixed a bug in FREEZE_THEN: the implementation using SUBGOAL_THEN was presuming two subgoals afterwards, yet it's possible that the ttac will solve the original goal completely and just leave the trivial one. This actually came up when I used FREEZE_THEN(fun t -> SIMP_TAC[t]) in order to fix some cases where things coincidentally worked thanks to the old PART_MATCH behaviour. Tue 23rd Sep 2008 tactics.ml I had compensated for the modified PART_MATCH and its consequent effect on MATCH_MP_TAC by tweaking a few proofs. But it was getting a bit tedious, and it didn't actually seem that the new behaviour was beneficial in this context since the invented variables are effectively frozen and so almost certainly not what's wanted. It seemed surprisingly difficult to put a basic wrapper round the existing PART_MATCH (I wanted some FREEZE-type thing, but you need to watch out for type instantiation and theorems with hypotheses). So I basically just copied the old PART_MATCH into this code. Mon 22nd Sep 2008 drule.ml Changed PART_MATCH so that it renames even free variables in the theorem if necessary when they would otherwise be uninstantiated (because they don't occur in the bit actually being matched) and clash with some post-instantiation variables. This is exemplified by the following: let th = TAUT `(g = g') ==> (g' ==> (t = t')) ==> (~g' ==> (e = e')) ==> ((if g then t else e) = (if g' then t' else e'))`;; let tm = `if CONS h' t' = [] then a else CONS h' t'`;; PART_MATCH (lhs o funpow 3 rand) th tm;; which used to give |- (CONS h' t' = [] <=> g') ==> (g' ==> a = t') ==> (~g' ==> CONS h' t' = e') ==> (if CONS h' t' = [] then a else CONS h' t') = (if g' then t' else e') but now gives |- (CONS h' t' = [] <=> g') ==> (g' ==> a = t'') ==> (~g' ==> CONS h' t' = e') ==> (if CONS h' t' = [] then a else CONS h' t') = (if g' then t'' else e') This was originally found by Jesse Bingham, who encountered a case where SIMP_CONV returned the "wrong" left-hand side thanks to this congruence rule instantiation. Tue 19th Aug 2008 lib.ml Removed a stray ";" in the definition of "mem'", pointed out by Mark Adams. For some reason OCaml didn't complain. Mon 11th Aug 2008 Boyer_Moore/* [new directory] Added a new directory containing code for Boyer-Moore automation. This was written by Petros Papapanagiotou, working with Jacques Fleuriot and starting from Richard Boulton's HOL88 code. Petros's thesis describing the setup is available at "http://www.inf.ed.ac.uk/publications/thesis/online/IM070466.pdf". Mon 11th Aug 2008 Examples/iter.ml [new file] Added a new file with theorems about the function iterator, from Marco et al. Occasionally I had used something similar but called it FUNPOW. Mon 11th Aug 2008 sets.ml Added SING_SUBSET = |- !s x. {x} SUBSET s <=> x IN s which was used in Marco et al's Cartan proof. Fri 8th Aug 2008 arith.ml, int.ml Changed the natural number and integer modulus operations so that n MOD 0 = n and n rem 0 = n. This was with the idea of removing the ~(n = 0) hypothesis from (a == b) (mod n) <=> (a MOD n = b MOD n) and others, though I didn't actually do that now. Thu 7th Aug 2008 int.ml Added INT_DIV_CONV and INT_REM_CONV for explicit calculations, and rolled these into INT_RED_CONV and INT_REDUCE_CONV. I should do something for INT_ARITH at some stage too. Thu 7th Aug 2008 int.ml Added definitions of "div" and "rem", using what Raymond Boute calls the "Euclidean" definition. Also added some supporting theorems, so we have: INT_WOP = |- (?x. &0 <= x /\ P x) <=> (?x. &0 <= x /\ P x /\ (!y. &0 <= y /\ P y ==> x <= y)) INT_DIVMOD_EXIST_0 = |- !m n. ?q r. if n = &0 then q = &0 /\ r = &0 else &0 <= r /\ r < abs n /\ m = q * n + r INT_DIVISION_0 = |- !m n. if n = &0 then m div n = &0 /\ m rem n = &0 else &0 <= m rem n /\ m rem n < abs n /\ m = m div n * n + m rem n INT_DIVISION = |- !m n. ~(n = &0) ==> m = m div n * n + m rem n /\ &0 <= m rem n /\ m rem n < abs n INT_DIVMOD_UNIQ = |- !m n q r. m = q * n + r /\ &0 <= r /\ r < abs n ==> m div n = q /\ m rem n = r Thu 7th Aug 2008 int.ml Added more forgotten analogues of real theorems, including the most recent set just below but also with one other strange omission: INT_LE_RMUL = |- !x y z. x <= y /\ &0 <= z ==> x * z <= y * z INT_POW_LE2_REV = |- !n x y. ~(n = 0) /\ &0 <= y /\ x pow n <= y pow n ==> x <= y INT_POW_LT2_REV = |- !n x y. &0 <= y /\ x pow n < y pow n ==> x < y INT_POW_EQ = |- !n x y. ~(n = 0) /\ &0 <= x /\ &0 <= y /\ x pow n = y pow n ==> x = y INT_POW_EQ_ABS = |- !n x y. ~(n = 0) /\ x pow n = y pow n ==> abs x = abs y Tue 22nd Apr 2008 real.ml Added a few theorems that seem like glaring omissions, even though they're not really hard to derive when needed: REAL_POW_LE2_REV = |- !n x y. ~(n = 0) /\ &0 <= y /\ x pow n <= y pow n ==> x <= y REAL_POW_LT2_REV = |- !n x y. &0 <= y /\ x pow n < y pow n ==> x < y REAL_POW_EQ = |- !n x y. ~(n = 0) /\ &0 <= x /\ &0 <= y /\ x pow n = y pow n ==> x = y REAL_POW_EQ_ABS = |- !n x y. ~(n = 0) /\ x pow n = y pow n ==> abs x = abs y Mon 21st Apr 2008 lib.ml Made a few small optimizations to the FPF code, basically inlining "ldb" and "is_undefined" to reduce function call overhead. Fri 29th Feb 2008 arith.ml Yet more simplifications suggested by Mark Bouler, though not using the same techniques, just more efficiently exploiting DIVMOD_UNIQ. Now the pairs DIV_0/MOD_0, DIV_1/MOD_1 and DIV_MULT/MOD_MULT are proved together using that and split into pairs. Tue 26th Feb 2008 arith.ml Simplified several proofs, this time of DIV_MOD, DIV_MONO, DIV_MOD_LT and DIV_LE_EXCLUSION_EXCLUSION using Mark Bouler's method. The only price was that the proof of the core lemma LE_RDIV_EQ became slightly longer so as not to rely on DIV_MONO. Also took the chance to slightly reorder things to separate the second suite of theorems about EXP from those about DIV and MOD. Sat 23rd Feb 2008 arith.ml Greatly simplified the proof of DIV_DIV using a trick that Mark Bouler showed me yesterday, using indirect equality (easy by MESON) and then the "Galois connection" lemma LE_RDIV_EQ. I should maybe see if this is more widely applicable. Tue 19th Feb 2008 arith.ml I always seem to be inventing ad-hoc lemmas to shuffle between different formulations of a natural number being strictly positive, so I added the following. Of course it will somewhat expand the search space in the simplifier but cuts through all the mess by having implications between all 6 pairs of formulations: LE_1 = |- (!n. ~(n = 0) ==> 0 < n) /\ (!n. ~(n = 0) ==> 1 <= n) /\ (!n. 0 < n ==> ~(n = 0)) /\ (!n. 0 < n ==> 1 <= n) /\ (!n. 1 <= n ==> 0 < n) /\ (!n. 1 <= n ==> ~(n = 0)) Sun 17th Feb 2008 iter.ml Slighly weakened hypothesis inequalities, formerly "m <= p" and "p <= n", in NUMSEG_COMBINE_R = |- !m p n. m <= p + 1 /\ p <= n ==> ((m..p) UNION ((p+1)..n) = m..n) NUMSEG_COMBINE_L = |- !m p n. m <= p /\ p <= n + 1 ==> ((m..(p-1)) UNION (p..n) = m..n) and added two combinations with sums: SUM_COMBINE_L = |- !f m n p. 0 < n /\ m <= n /\ n <= p + 1 ==> sum (m..n - 1) f + sum (n..p) f = sum (m..p) f SUM_COMBINE_R = |- !f m n p. m <= n + 1 /\ n <= p ==> sum (m..n) f + sum (n + 1..p) f = sum (m..p) f and an unrelated but useful theorem: SUM_ABS_LE = |- !f g s. FINITE s /\ (!x. x IN s ==> abs (f x) <= g x) ==> abs (sum s f) <= sum s g Wed 13th Feb 2008 iter.ml Added two forms of partial summation: SUM_PARTIAL_SUC = |- !f g m n. sum (m..n) (\k. f k * (g(k + 1) - g k)) = (if m <= n then f(n + 1) * g(n + 1) - f m * g m - sum (m..n) (\k. g(k + 1) * (f(k + 1) - f k)) else &0) SUM_PARTIAL_PRE = |- !f g m n. sum (m..n) (\k. f k * (g k - g(k - 1))) = (if m <= n then f(n + 1) * g n - f m * g(m - 1) - sum (m..n) (\k. g k * (f(k + 1) - f k)) else &0) Tue 12th Feb 2008 itab.ml So it turns out that my tweak on 31st Jan last year was not right after all. Sean McLaughlin pointed out that the following now didn't work: ITAUT `(((~ ~ (p \/ ~p)) ==> (p \/ ~p)) ==> (p \/ ~p))`;; So I just went back to the old code in this respect, putting back FIRST_ASSUM instead of FIRST_X_ASSUM. Tue 5th Feb 2008 real.ml Added: REAL_EQ_INV2 = |- !x y. inv x = inv y <=> x = y REAL_INV_EQ_1 = |- !x. inv x = &1 <=> x = &1 Tue 5th Feb 2008 arith.ml Added a bunch of theorems that are an obvious gap in the arithmetic suite; several variants of these are in Examples/prime.ml, but often decorated with "SUC". So I added "SUC" suffixes to those versions. EXP_EQ_1 = |- !x n. x EXP n = 1 <=> x = 1 \/ n = 0 EXP_MONO_LE_IMP = |- !x y n. x <= y ==> x EXP n <= y EXP n EXP_MONO_LT_IMP = |- !x y n. x < y /\ ~(n = 0) ==> x EXP n < y EXP n EXP_MONO_LE = |- !x y n. x EXP n <= y EXP n <=> x <= y \/ n = 0 EXP_MONO_LT = |- !x y n. x EXP n < y EXP n <=> x < y /\ ~(n = 0) EXP_MONO_EQ = |- !x y n. x EXP n = y EXP n <=> x = y \/ n = 0 Mon 4th Feb 2008 calc_rat.ml, int.ml Added something I should probably have done a long time ago: ASM_ARITH_TAC ASM_INT_ARITH_TAC ASM_REAL_ARITH_TAC These basically pop every assumption that isn't universally quantified and use them as additional hypotheses. I was getting sick of using repeated idioms like "REPEAT(POP_ASSUM MP_TAC) THEN ARITH_TAC". Mon 4th Feb 2008 arith.ml Added MULT_DIV_LE = |- !m n p. ~(p = 0) ==> m * (n DIV p) <= (m * n) DIV p Sun 3rd Feb 2008 arith.ml Added EQ_EXP = |- !x m n. x EXP m = x EXP n <=> (if x = 0 then m = 0 <=> n = 0 else x = 1 \/ m = n) Tue 29th Jan 2008 sets.ml Added: FINITE_FINITE_PREIMAGE_GENERAL = |- !f s t. FINITE t /\ (!y. y IN t ==> FINITE {x | x IN s /\ f x = y}) ==> FINITE {x | x IN s /\ f x IN t} FINITE_FINITE_PREIMAGE = |- !f t. FINITE t /\ (!y. y IN t ==> FINITE {x | f x = y}) ==> FINITE {x | f x IN t} Thu 24th Jan 2008 Unity/make.ml, Unity/mk_unity_prog.ml Made a fix indicated by Flemming to the "mk_unity_prog.ml" file so that now it loads unproblematically, and hence uncommented its load in the "make.ml" root. Fri 18th Jan 2008 Unity [new directory] Added the first version of Flemming Andersen's Unity theory, which he has ported to HOL Light. Wed 16th Jan 2008 real.ml Added REAL_POW_ZERO = |- !n. &0 pow n = (if n = 0 then &1 else &0) REAL_POW_MONO_INV = |- !m n x. &0 <= x /\ x <= &1 /\ n <= m ==> x pow m <= x pow n Tue 15th Jan 2008 arith.ml Finally added FACT_NZ = |- !n. ~(FACT n = 0). Of course it's trivial to generate, but I seem to be doing it quite often. Wed 7th Jan 2008 hol.ml Changed (back?) from Filename.temp_dir to "/tmp", since I just discovered on getting a new laptop that the OCaml in Cygwin, still 3.08, doesn't have that. It seems a pity to force people to upgrade their OCaml just for this slightly more elegant solution. Maybe in a year or two I'll change it back. Wed 2nd Jan 2008 define.ml Put in a rewritten version of this stuff, with a few new proforma theorems and a somewhat different (simpler?) method for hacking the proformas to include many variables. Most notably, this adds support for the new "match" and "function" constructs, but has some other enhancements too such as "MAP". Also the "instantiate_casewise_recursion" folds in the tupling of multiple arguments, which was otherwise only done in functions further out. Sat 29th Dec 2007 Permutation/*.ml Installed new version of the Permutation material from Marco Maggesi. Thu 27th Dec 2007 sets.ml Added SIMPLE_IMAGE_GEN = |- !f p. {f x | P x} = IMAGE f {x | P x} Sun 23rd Dec 2007 iter.ml Added SUM_IMAGE_LE = |- !f g s. FINITE s /\ (!x. x IN s ==> &0 <= g(f x)) ==> sum (IMAGE f s) g <= sum s (g o f) and IN_NUMSEG_0 = |- !m n. m IN 0..n <=> m <= n Sun 23rd Dec 2007 sets.ml Added EXISTS_IN_UNIONS = |- !P s. (?x. x IN UNIONS s /\ P x) <=> (?t x. t IN s /\ x IN t /\ P x) Tue 18th Dec 2007 realarith.ml Tweaked the "eliminate_construct" subfunction in GEN_REAL_ARITH to check that the term it's finding is actually free in the whole term. Previously, it could go into an infinite loop in cases like this by repeatedly attempting to replace non-free subterms: REAL_ARITH `&0 <= abs(sum s (\x. abs x))`;; Tue 18th Dec 2007 real.ml Added REAL_POW_1_LT = |- !n x. ~(n = 0) /\ &0 <= x /\ x < &1 ==> x pow n < &1 REAL_POW_LT_1 = |- !n x. ~(n = 0) /\ &1 < x ==> &1 < x pow n which I was a bit surprised not to have already, and also the integer counterparts INT_POW_1_LT and INT_POW_LT_1. Tue 18th Dec 2007 Examples/floor.ml To be more consistent with other names, redefined: REAL_EQ_INTEGERS = |- !x y. integer x /\ integer y ==> (x = y <=> abs (x - y) < &1) and renamed the old theorem that had that name: REAL_EQ_INTEGERS_IMP = |- !x y. integer x /\ integer y /\ abs (x - y) < &1 ==> x = y Sun 16th Dec 2007 cart.ml Added this finiteness theorem (which is actually quite tedious to prove!) FINITE_CART = |- !P. (!i. 1 <= i /\ i <= dimindex (:N) ==> FINITE {x | P i x}) ==> FINITE {v | !i. 1 <= i /\ i <= dimindex (:N) ==> P i (v$i)} Mon 10th Dec 2007 iter.ml Added two slighly generalized forms of [N]SUM_IMAGE_GEN: NSUM_GROUP = |- !f g s t. FINITE s /\ IMAGE f s SUBSET t ==> nsum t (\y. nsum {x | x IN s /\ f x = y} g) = nsum s g SUM_GROUP = |- !f g s t. FINITE s /\ IMAGE f s SUBSET t ==> sum t (\y. sum {x | x IN s /\ f x = y} g) = sum s g Sat 8th Dec 2007 sets.ml Added UNIONS_UNION = |- !s t. UNIONS (s UNION t) = UNIONS s UNION UNIONS t Tue 4th Dec 2007 iter.ml Added the following (maybe should also do an ITERATE version sometime). NSUM_UNIONS_NONZERO = |- !f s. FINITE s /\ (!t. t IN s ==> FINITE t) /\ (!t1 t2 x. t1 IN s /\ t2 IN s /\ ~(t1 = t2) /\ x IN t1 /\ x IN t2 ==> f x = 0) ==> nsum (UNIONS s) f = nsum s (\t. nsum t f) SUM_UNIONS_NONZERO = |- !f s. FINITE s /\ (!t. t IN s ==> FINITE t) /\ (!t1 t2 x. t1 IN s /\ t2 IN s /\ ~(t1 = t2) /\ x IN t1 /\ x IN t2 ==> f x = &0) ==> sum (UNIONS s) f = sum s (\t. sum t f) Tue 4th Dec 2007 sets.ml Added SUBSET_CARD_EQ = |- !s t. FINITE t /\ s SUBSET t ==> (CARD s = CARD t <=> s = t) Mon 3rd Dec 2007 wf.ml, define.ml, passim Changed "measure" to "MEASURE" (both the constant name and the ML binding of its definition). This frees up lowercase "measure" for Lebesgue measure, which I've finally started seriously working with. Sun 2nd Dec 2007 set.ml Added UNIONS_SUBSET = |- !f t. UNIONS f SUBSET t <=> (!s. s IN f ==> s SUBSET t) SUBSET_UNIONS = |- !f g. f SUBSET g ==> UNIONS f SUBSET UNIONS g Fri 23rd Nov 2007 iter.ml Added CARD_UNIONS = |- !s. FINITE s /\ (!t. t IN s ==> FINITE t) /\ (!t u. t IN s /\ u IN s /\ ~(t = u) ==> t INTER u = {}) ==> CARD(UNIONS s) = nsum s CARD Thu 22nd Nov 2007 sets.ml To celebrate Thanksgiving I proved that integrals of arbitrary functions over a set of measure zero are 0. I ended up using a lemma about sums, and I thought I'd add it to the core. In one sense it's a bit ad hoc, but corresponds to a very straightforward intuition: all elements of one sum are dominated by some of another, and the rest of the second sum are nonnegative. SUM_LE_INCLUDED = |- !f g s t i. FINITE s /\ FINITE t /\ (!y. y IN t ==> &0 <= g y) /\ (!x. x IN s ==> (?y. y IN t /\ i y = x /\ f x <= g y)) ==> sum s f <= sum t g Tue 20th Nov 2007 int.ml Added integer versions of some recent additions I'd forgotten, including the latest: INT_BOUNDS_LE = |- !x k. --k <= x /\ x <= k <=> abs x <= k INT_BOUNDS_LT = |- !x k. --k < x /\ x < k <=> abs x < k INT_MUL_POS_LT = |- !x y. &0 < x * y <=> &0 < x /\ &0 < y \/ x < &0 /\ y < &0 INT_MUL_POS_LE = |- !x y. &0 <= x * y <=> x = &0 \/ y = &0 \/ &0 < x /\ &0 < y \/ x < &0 /\ y < &0 INT_LE_MUL_EQ = |- (!x y. &0 < x ==> (&0 <= x * y <=> &0 <= y)) /\ (!x y. &0 < y ==> (&0 <= x * y <=> &0 <= x)) INT_LT_MUL_EQ = |- (!x y. &0 < x ==> (&0 < x * y <=> &0 < y)) /\ (!x y. &0 < y ==> (&0 < x * y <=> &0 < x)) To get the last two I had to add a little tweak to INT_OF_REAL_THM to handle conjunctive equations; previously it only worked if all the universal quantifiers were at the outside. Tue 20th Nov 2007 real.ml Added two theorems that I keep regenerating: ( REAL_BOUNDS_LE ) : thm = |- !x k. --k <= x /\ x <= k <=> abs x <= k ( REAL_BOUNDS_LT ) : thm = |- !x k. --k < x /\ x < k <=> abs x < k The symmetric form of the first one was already there as REAL_ABS_BOUNDS, but I've never actually used it (however, since it's used by Freek and Tom, I didn't delete it). Mon 19th Nov 2007 real.ml Added REAL_MUL_POS_LT = |- !x y. &0 < x * y <=> &0 < x /\ &0 < y \/ x < &0 /\ y < &0 REAL_MUL_POS_LE = |- !x y. &0 <= x * y <=> x = &0 \/ y = &0 \/ &0 < x /\ &0 < y \/ x < &0 /\ y < &0 Wed 14th Nov 2007 ind-types.ml Added a trivial tweak in "prove_cases_thm" so that it works on num_INDUCTION, where one of the "constructors" 0 is not really a constant. Just check for reflexivity in "prove_triv" subfunction. Wed 14th Nov 2007 calc_rat.ml Added a little extra feature to REAL_FIELD so that it uses any strict inequality hypotheses s < t to add a hypothesis ~(s = t). Previously, this only happened if there was a correlation in a branch with an explicit equation (s = t), e.g. introduced in the elimination of inverses. Now, for example, we solve this which failed before: REAL_FIELD `&0 < x ==> x * y = x * z ==> y = z`;; Tue 13th Nov 2007 sets.ml Added SUBSET_ANTISYM_EQ = |- !s t. s SUBSET t /\ t SUBSET s <=> s = t Tue 13th Nov 2007 Modified "pa_j_3.10.ml" for the latest version of camlp5 (5.02). Needed a bit of grubbing round; in particular I had to add "translate_operator" into the new "val_ident" clause. Mon 12th Nov 2007 Added FORALL_UNWIND_CONV; I wanted some such operation inside the new "define" wellfoundedness-guessing, and it seemed useful and general enough to do properly. Wed 7th Nov 2007 sets.ml Added SUBSET_INTER = |- !s t u. s SUBSET t INTER u <=> s SUBSET t /\ s SUBSET u Tue 6th Nov 2007 printer.ml Modified the printer so that when trying to print a "function" or "match" construct it *first* makes sure the construct is in the canonical form before printing anything. Otherwise the kind of more general cases that crop up in admissibility lemmas show up as an exception when printing. Tue 30th Oct 2007 Makefile Changed "[ ... ]" to "test", and "==" to "=" in the Makefile, which for some reason barfed on my new Ubuntu box but not on other platforms. Anyway it seems more correct/portable. Wed 17th Oct 2007 pair.ml, ind-types.ml Changed the definitions of _MATCH and _FUNCTION from |- _MATCH = \e r. (@) (r e) |- _FUNCTION = \r x. (@) (r x) to |- _MATCH = \e r. if (?!) (r e) then (@) (r e) else @z. F |- _FUNCTION = \r x. if (?!) (r x) then (@) (r x) else @z. F I'd discovered painfully that the former does not interact well with the special treatment of tail recursion in the "define.ml" setup, because in general there is some nondeterminacy there. The new definition clears up this problem, and is also perhaps a little more satisfying on general grounds anyway, because it avoids any "accidental" assignments in cases where the match isn't really well-defined, and so provides a bit more automatic debugging. The only changes necessary were to the newish MATCH_CONV. They were fairly non-invasive and should only have a small efficiency cost (the expensive unwinding etc. is only done once even though the predicate appears twice in the underlying expression). Tue 16th Oct 2007 sets.ml Added the infinitude of the new string type: string_INFINITE = |- INFINITE (:(char)list) Fri 12th Oct 2007 printer.ml, parser.ml Finished the job, more or less, by adding printer support for strings too, and also improving the treatment of escapes so they correspond to OCaml (or at least to the OCaml manual). That should more or less give a parse/print equivalence. Fri 12th Oct 2007 list.ml, parser.ml Added a new type "char" defined like the old HOL88 type "ascii" with a constructor ASCII taking 8 Booleans, together with the abbreviation of "string" for "char list" (not a separate type, which aids orthogonality). Added support for strings in the parser. I should probably do a more systematic set of escape sequences; as it is I only have \n, \\ and \". Mon 8th Oct 2007 Examples/sos.ml I'd noticed that my REAL_SOSFIELD was in step with an earlier mistaken optimization of REAL_FIELD, which generated fewer cases but was not in general complete. It probably doesn't matter, but now it's fixed. Mon 8th Oct 2007 Examples/sos.ml I wanted to take a careful look at the iterative deepening I was doing in SOS w.r.t. strict inequalities to make absolutely sure that it was indeed theoretically complete. It does seem all right in principle, though I can think of other state space explorations that might be more efficient to try. But anyway I did spot one little error in the case where there are no strict inequalities in the hypotheses. Changed maximal degree of the polynomial to consider from 0 to 1, since otherwise one is simply duplicating work. From let k = if e = 0 then 1 else d / e in to let k = if e = 0 then 0 else d / e in Tue 11th Sep 2007 basics.ml Removed duplicate definitions of "dest_forall" and "strip_forall" inside the "dest_gabs" function. These were, I guess, a relic of the time when these functions too were defined in "bool.ml", as "mk_forall" still is. This was pointed out by Viorel Preoteasa. Tue 11th Sep 2007 ind-types.ml Added UNWIND_CONV (which was useful as a subroutine in the following, and seems useful enough to expose) and MATCH_CONV to reduce matches applied to specific cases. Also added the "simple" case of MATCH_CONV, where an unambiguous reduction is achieved automatically, to the default conversions. Sat 8th Sep 2007 ind-defs.ml, sets.ml, recursion.ml Added a list "the_inductive_definitions" and associated benignity checking. This was a reaction to the observation by Norbert Voelker that reloading "sets.ml" gives an error because of the inductive definition of FINITE. (For extra credit, I could search for redefinition in exactly the same form with variables in the input, in case it's generated other than by the quotation parser.) However, I had also to add a type constraint to make sure the types in the empty and insert claises really are the same and hence it is an instance. I also discovered I still need benignity checking in "recursion.ml" to handle FINREC. So I added that too, though didn't expose the theorem list in this case. And once again I needed to add a type constrain (many from "list.ml" won't work because the types are insufficiently constrained). All a bit crude and inadequate, but it does solve the top-level problem: "sets.ml" *can* now be reloaded. Fri 7th Sep 2007 Examples/prover9.ml [new file] Added to the standard release the simple prover9 interface I hacked up at Marktoberdorf and slightly refined more recently. It could still use more work (e.g. a disjunctive splitter, and more systematic treatment of "pure equality islands") but it's nice to have. Fri 7th Sep 2007 Examples/update_database.ml Switched the load order of "env.cmo" and "clflags.cmo", so the latter now comes first. This seems to be necessary for 3.10.0? Fri 7th Sep 2007 Examples/pratt.ml, Examples/pocklington.ml Renamed the image name just "gp" not "parigp", which seems to be what it (now) is OOTB. Fri 7th Sep 2007 make.ml Added (and documented) a "checkpoint" function to create a checkpoint with CryoPID. This is a temporary compromise since I have trouble getting it to work cleanly in a batch build. Wed 5th Sep 2007 hol.ml, passim Changed "temp_path" to read the OCaml Filename.temp_dir_name at the outset, though the user can set it. Also modified a few bits and pieces to use this interface consistently (e.g. SOS and Minisat) where previously there was a hardwired "/tmp" pathname. Tue 4th Sep 2007 preterm.ml Added some error traps to produce a "Typechecking error" report rather than the less helpful "not found". Mon 3rd Sep 2007 tactics.ml Put in a fix for an assumption of equality instead of alpha-convertibility of hypotheses into VALID, which indirectly affects "e". Previously, for example, if a goal had an assumption `!x. x = 1`, then using ASSUME(`!z. z = 1`) would give a validty failure. Possibly a more systematic approach would be to replace each "union" with "term_union", but I wanted to be minimally invasive. Mon 3rd Sep 2007 printer.ml Fixed something pointed out by Carl Witty a while ago: a construct `DECIMAL mmm nnn` where nnn is a numeral but not a power of 10 prints misleadingly, e.g. `DECIMAL 99 77` as `#.19`. Mon 3rd Sep 2007 preterm.ml, parser.ml Added a warning if something starting with a digit is not a valid numeral, e.g. `0xgh` or `0b12`. It seems unintuitive to parse this as a variable without any warning. To support this change I moved the character discrimination functions back from "parser.ml" to "preterm.ml". Mon 3rd Sep 2007 bool.ml Recoded IMP_TRANS and EQ_IMP_RULE with more elaborate versions that correspond exactly (I hope) to their specification, taking the union of the assumption list even if some of the implication components are in the assumptions; with the old implementations one could sometimes get fewer assumptions. The case of IMP_TRANS was pointed out on info-hol by Tony Johnson (complaining about HOL4, but it also applies to HOL Light), and I noticed a similar problem with EQ_IMP_RULE when I dived in to fix that. Mon 3rd Sep 2007 printer.ml Modified the printer so it prints parentheses exactly round generalized varstructs and not regular ones. This fixes one flat bug where an iterated mixture of generalized and non-generalized abstractions would print without the parentheses necessary to re-parse the same thing, e.g. `\x y,z. bod` instead of `\x (y,z). bod`. It also seems nice to have a visual cue that generalized and non-generalized abstractions differ even in the unit variable case. I would have liked to establish the same thing on input, but that's not completely trivial given that for a long time things exist as preterms and only get mapped to any sort of abstraction later on. Fri 31st Aug 2007 tactics.ml Fixed a bug in the printing of goalstacks where the boxes inside hypotheses were getting offset by 2, e.g. 1 [`i IN 1 .. dimindex (:?162964) /\ j IN 1 .. dimindex (:?162964) /\ ~(i = j)`] It turns out I'd misunderstood what "print_as" does: it doesn't print a string in a fixed field but just sets the counts as if it did. So I added my own custom function to print the assumption number padded up to 3 digits if it would otherwise be shorter. Fri 31st Aug 2007 make.ml Thanks to an email from Vic Zandy, I finally solved the problem of incorrect PIDs on recent Linuxes. The issue seems to be that the glibc implementation of getpid is cacheing the PID in userspace, so all you have to do is use some other method of getting the PID. Vic suggested reading /proc/self/stat, but for my application it's even easier to look at $PPID. Also, experimenting on Knoppix I found that adding a 1-second sleep before the kill call deals with the intermittent race issue. I don't really understand this but I put it in anyway. Thu 30th Aug 2007 basics.ml, equal.ml Changed "mk_primed_var" so that it ignores hidden constants. This entailed moving it forward a bit till after the "hide" stuff, and therefore I moved it from "basics.ml" to "equal.ml". Thu 30th Aug 2007 lib.ml, grobner.ml, help.ml, meson.ml, tactics.ml A suggestion from Jesse Alama was a quiet loading mode. When experimenting with this I noticed quite a few places where I used raw "print_string" (and occasionally "print_newline" and maybe others) instead of their Format versions, so fixed that for consistency. Tue 28th Aug 2007 parser.ml Fixed a bug in pfrees where "can int_of_string" actually built in a size limitation so large numerals would be treated as variables! Now uses "num_of_string" instead. Tue 28th Aug 2007 parser.ml, pair.ml, printer.ml Added OCaml-like pattern-matching constructs to the parser, together with definitions in "pair.ml" for the supporting constants _SEQPATTERN, _UNGUARDED_PATTERN, _GUARDED_PATTERN, _MATCH and _FUNCTION. The main incompatible change is making the following reserved words: match with function -> when Tue 28th Aug 2007 parser.ml Modified the "nocommapreterm" so it won't fail if "," doesn't have infix parse status (just a trivial little bug I noticed). Tue 28th Aug 2007 realarith.ml Changed some quotations in GEN_LINEAR_PROVER to force them to evaluate here. Although they get partially evaluated anyway when this is applied to the first argument (like LINEAR_PROVER), that still introduces a nasty dependency on the overload prioritization, and I hit this while checking Lars's more general code, discovering that with prioritize_int() set even "GEN_REAL_ARITH REAL_LINEAR_PROVER tm" fails. Mon 27th Aug 2007 meson.ml Removed the default "time" inside the core function so it doesn't automatically report CPU times; I felt the output was already quite verbose enough. Fri 24th Aug 2007 int.ml Fixed bug of using int_sub instead of INT_SUB in the definition of INT_RED_CONV, which was causing evaluation of integer constant expressions to fail. This and the last item while getting my car serviced. Fri 24th Aug 2007 Examples/integer.ml [new file] Put in a convenient place the basic development of divisibility properties over Z. Much of it is automatic with INTEGER_TAC but some slightly tedious hacking is needed for signs of GCDs and suchlike. Fri 24th Aug 2007 fusion.ml Added back an internal list of definitions (called "the_definitions", but not exported and distinct from the same thing defined later) and a function "definitions" to get at them. Mark Adams pointed out that otherwise you can't distinguish axioms that should be sound from those that aren't because they clash with definitions. Thu 23rd Aug 2007 sets.ml Added two theorems that I can hardly believe I'd never wanted before: CARD_DIFF = |- !s t. FINITE s /\ t SUBSET s ==> CARD(s DIFF t) = CARD s - CARD t HAS_SIZE_DIFF = |- !s t m n. s HAS_SIZE m /\ t HAS_SIZE n /\ t SUBSET s ==> s DIFF t HAS_SIZE m - n Sun 29th Jul 2007 Makefile, hol.ml Modified things so that camlp5 more or less automatically gets used if the OCaml version is >= 3.10. Also reduced the proliferation of pa_j files by employing "cut" on the version number to look at the major number only. Wed 25th Jul 2007 real.ml, int.ml Added REAL_LT_SQUARE_ABS = |- !x y. abs x < abs y <=> x pow 2 < y pow 2 and INT_LT_SQUARE_ABS = |- !x y. abs x < abs y <=> x pow 2 < y pow 2 Mon 23rd Jul 2007 iter.ml, passim Added a theorem from Lars Schewe, though this time with a different proof using SUM_POS_EQ_0. SUM_ZERO_EXISTS = |- !u s. FINITE s /\ sum s u = &0 ==> (!i. i IN s ==> u i = &0) \/ (?j k. j IN s /\ u j < &0 /\ k IN s /\ u k > &0) Added Lars to copyright for this file and took the chance to update all the others from "2006" to "2007". Mon 23rd Jul 2007 hol.ml and a few others Changed HOLDIR to HOLLIGHT_DIR at Hasan's request, to make it easier to work in a multi-HOL environment. Mon 23rd Jul 2007 Minisat/zc2mso [new directory] Added Hasan's new "zc2mso" translator, the C++ source and README. Also slightly modified the main README to point out its existence. Fri 13th Jul 2007 real.ml Added REAL_LE_MUL_EQ = |- (!x y. &0 < x ==> (&0 <= x * y <=> &0 <= y)) /\ (!x y. &0 < y ==> (&0 <= x * y <=> &0 <= x)) REAL_LT_MUL_EQ = |- (!x y. &0 < x ==> (&0 < x * y <=> &0 < y)) /\ (!x y. &0 < y ==> (&0 < x * y <=> &0 < x)) Wed 11th Jul 2007 sets.ml Fixed FINITE_RESTRICT to |- !s P. FINITE s ==> FINITE {x | x IN s /\ P x} instead of vacuous quantifeer |- !s p. FINITE s ==> FINITE {x | x IN s /\ P x} as pointed out by Lars Schewe. Wed 11th Jul 2007 iter.ml Added some new theorems about sums, based on Lars Schewe's files, and using some of the same proofs modulo trivial changes. First of all, renamed the old ITERATE_CASES -> ITERATE_EXPAND_CASES in order to make the more natural namespace available. Then added ITERATE_INCL_EXCL = |- !op. monoidal op ==> (!s t f. FINITE s /\ FINITE t ==> op (iterate op s f) (iterate op t f) = op (iterate op (s UNION t) f) (iterate op (s INTER t) f)) ITERATE_CASES = |- !op. monoidal op ==> (!s P f g. FINITE s ==> iterate op s (\x. if P x then f x else g x) = op (iterate op {x | x IN s /\ P x} f) (iterate op {x | x IN s /\ ~P x} g)) as well as versions for natural number sums NSUM_INCL_EXCL = |- !s t f. FINITE s /\ FINITE t ==> nsum s f + nsum t f = nsum (s UNION t) f + nsum (s INTER t) f NSUM_CASES = |- !s P f g. FINITE s ==> nsum s (\x. if P x then f x else g x) = nsum {x | x IN s /\ P x} f + nsum {x | x IN s /\ ~P x} g and real sums: SUM_INCL_EXCL = |- !s t f. FINITE s /\ FINITE t ==> sum s f + sum t f = sum (s UNION t) f + sum (s INTER t) f SUM_CASES = |- !s P f g. FINITE s ==> sum s (\x. if P x then f x else g x) = sum {x | x IN s /\ P x} f + sum {x | x IN s /\ ~P x} g Wed 11th Jul 2007 calc_rat.ml, Complex/complex.ml Undid the change to the two FIELD rules made on 23rd Feb 06, though keeping the two later tweaks (inverse-inverse and checking freeness). The problem is that sometimes the "inefficient" case split is necessary to pick up correlations between nonzeroness assumptions. I thought this was more important than efficiency. For example, the following previously failed: REAL_FIELD `~(c = &0) /\ ~(c' = &0) /\ ~(c * c' - s * s' = &0) ==> (s * c' + c * s') / (c * c' - s * s') = (s / c + s' / c') / (&1 - s / c * s' / c')`;; even though both the following worked: REAL_FIELD `~(c = &0) /\ ~(c' = &0) /\ ~(c * c' - s * s' = &0) /\ ~(&1 - s / c * s' / c' = &0) ==> (s * c' + c * s') / (c * c' - s * s') = (s / c + s' / c') / (&1 - s / c * s' / c')`;; REAL_FIELD `~(c = &0) /\ ~(c' = &0) /\ ~(c * c' - s * s' = &0) ==> ~(&1 - s / c * s' / c' = &0)`;; Now they all work, albeit a bit more slowly. Wed 23rd May 2007 Multivariate/vectors.ml Installed NORM_ARITH in the main file and already used it in several places in place of explicit proofs. Tue 22nd May 2007 Examples/sos.ml Added a little wrapper REAL_NONLINEAR_SUBST_PROVER round REAL_NONLINEAR_PROVER, which tries to first substitute variables in equations. This must almost invariably be a good idea. The particular motivation was the application of Solovay's procedure to the triangle law, which gives rise to the real problem: |- &0 <= c /\ &0 <= z /\ z pow 2 = h pow 2 * d + c /\ &0 <= y /\ y pow 2 = d /\ &0 <= x /\ x pow 2 = d + &2 * h * d + h pow 2 * d + c /\ y + z < x ==> F With the substitution this takes about 18 seconds; without it about 180. So here it's worthwhile, and I guess this is true more generally. Tue 1st May 2007 simp.ml Well, backed off that change for a while, since it broke things in Jordan/metric_spaces.ml and I ended up not really using it myself. Tue 24th Apr 2007 simp.ml Modified ONCE_DEPTH_SQCONV and ONCE_SIMPLIFY_CONV so that the prover list in the subcalls is modified to try assumption if all else fails. Then modified ONCE_SIMP_TAC (and so implicitly ONCE_ASM_SIMP_TAC) so that they handle the possible new assumption and split it off as a subgoal. The net effect, I hope, is that ONCE_SIMP_TAC will now actually be useful for cases where you really want to split off the lemma. Tue 13th Mar 2007 iter.ml Added SUM_DELETE_CASES = |- !f s a. FINITE s ==> sum (s DELETE a) f = (if a IN s then sum s f - f a else sum s f) Tue 13th Feb 2007 Added a theorem that only recently occurred to me; it can occasionally be useful not to have to discharge a nonnegativity assumption when the exponent is odd. REAL_POW_LE2_ODD = |- !n x y. x <= y /\ ODD n ==> x pow n <= y pow n Added the counterpart for the integers and also some other theorems I seem to have forgotten to transfer: INT_LE_SQUARE_ABS = |- !x y. abs x <= abs y <=> x pow 2 <= y pow 2 INT_SOS_EQ_0 = |- !x y. x pow 2 + y pow 2 = &0 <=> x = &0 /\ y = &0 INT_POW_LE2_ODD = |- !n x y. x <= y /\ ODD n ==> x pow n <= y pow n Sat 3rd Feb 2007 tactics.ml Realized I'd put the tweaked error message only in a place where failure will occur if there is a current goal but it has no subgoals. Reworded that one slightly and added it where I should have, in "refine". Thu 1st Feb 2007 iter.ml Added SUM_DIFFS = |- !m n. sum (m..n) (\k. f k - f (k + 1)) = (if m <= n then f m - f (n + 1) else &0) Wed 31st Jan 2007 itab.ml Made a small tweak to the G3 implementation, which I noticed when reminding myself about stuff in order to answer a question on hol-info. In the left-implication rule, I was using FIRST_ASSUM, but it should/could be FIRST_X_ASSUM, I guess. Wed 31st Jan 2007 Makefile Added "hol.complex" to the Makefile, since I'm using it heavily at the moment, and it's becoming non-trivial. Wed 31st Jan 2007 parser.ml, tactics.ml Changed some annoying error messages, the "unexpected junk" to "unparsed input" and also a more informative error than "hd" when trying to apply a tactic with no goals. Mon 22nd Jan 2007 simp.ml Fixed a plain bug in ABBREV_TAC which failed when abbreviations had more than one argument due to a trivial blunder. Sun 21st Jan 2007 sets.ml Added a suite of theorems to capture a pattern of reasoning I seem to go through very often: TRANSITIVE_STEPWISE_LE = |- !R. (!x. R x x) /\ (!x y z. R x y /\ R y z ==> R x z) /\ (!n. R n (SUC n)) ==> (!m n. m <= n ==> R m n) TRANSITIVE_STEPWISE_LT = |- !R. (!x y z. R x y /\ R y z ==> R x z) /\ (!n. R n (SUC n)) ==> (!m n. m < n ==> R m n) TRANSITIVE_STEPWISE_LE_EQ = |- !R. (!x. R x x) /\ (!x y z. R x y /\ R y z ==> R x z) ==> ((!m n. m <= n ==> R m n) <=> (!n. R n (SUC n))) TRANSITIVE_STEPWISE_LT_EQ = |- !R. (!x y z. R x y /\ R y z ==> R x z) ==> ((!m n. m < n ==> R m n) <=> (!n. R n (SUC n))) Sat 20th Jan 2007 sets.ml Added IMAGE_UNIONS = |- !f s. IMAGE f (UNIONS s) = UNIONS (IMAGE (IMAGE f) s) Tue 16th Jan 2007 iter.ml Added four new theorems relevant to iterated operations: MONOIDAL_AC = |- !op. monoidal op ==> (!a. op (neutral op) a = a) /\ (!a. op a (neutral op) = a) /\ (!a b. op a b = op b a) /\ (!a b c. op (op a b) c = op a (op b c)) /\ (!a b c. op a (op b c) = op b (op a c)) ITERATE_OP = |- !op. monoidal op ==> (!f g s. FINITE s ==> iterate op s (\x. op (f x) (g x)) = op (iterate op s f) (iterate op s g)) ITERATE_SUPERSET = |- !op. monoidal op ==> (!f u v. u SUBSET v /\ (!x. x IN v /\ ~(x IN u) ==> f x = neutral op) ==> iterate op v f = iterate op u f) ITERATE_IMAGE_NONZERO = |- !op. monoidal op ==> (!g f s. FINITE s /\ (!x y. x IN s /\ y IN s /\ ~(x = y) /\ f x = f y ==> g (f x) = neutral op) ==> iterate op (IMAGE f s) g = iterate op s (g o f)) The penultimate one made me realize that the finiteness condition on NSUM_SUPERSET and SUM_SUPERSET was superfluous, so I removed it. Mon 15th Jan 07 pair.ml, wf.ml, Examples/hol88.ml, passim Retired PAIRED_BETA_CONV, moving it from "pair.ml" just into the HOL88 compatibility files. In a handful of modern proofs scrubbed it when it's now subsumed by rewrites; in older files mechanically changed to "GEN_BETA_CONV". In the process, I've quietly generalized let_CONV to work over other varstructs subject to analogous restrictions to the pair instance. Mon 15th Jan 07 class.ml, trivia.ml, pair.ml, num.ml, ind-types.ml Changed "inductive_type_store" to be built up one at a time as the new types are added, really just so that GEN_BETA_CONV on pairs already works inside the pairs file plus a bit of intellectual consistency. Mon 15th Jan 07 pair.ml Added LAMBDA_PAIR_THM = |- (\p. P p) = (\(x,y). P (x,y)) Mon 15th Jan 07 iter.ml Added the following suite, which are just tedious enough not to want to derive by hand when needed: ITERATE_UNION_NONZERO = |- !op. monoidal op ==> (!f s t. FINITE s /\ FINITE t /\ (!x. x IN s INTER t ==> f x = neutral op) ==> iterate op (s UNION t) f = op (iterate op s f) (iterate op t f)) NSUM_UNION_NONZERO = |- !f s t. FINITE s /\ FINITE t /\ (!x. x IN s INTER t ==> f x = 0) ==> nsum (s UNION t) f = nsum s f + nsum t f SUM_UNION_NONZERO = |- !f s t. FINITE s /\ FINITE t /\ (!x. x IN s INTER t ==> f x = &0) ==> sum (s UNION t) f = sum s f + sum t f Sat 13th Jan 07 iter.ml Added the following; I already have the vector version, and I didn't yet decide to add the natural number version. Should really do a generic one. SUM_IMAGE_NONZERO = |- !d i s. FINITE s /\ (!x y. x IN s /\ y IN s /\ ~(x = y) /\ i x = i y ==> d (i x) = &0) ==> sum (IMAGE i s) d = sum s (d o i) Thu 11th Jan 07 sets.ml Added another theorem: INTER_UNIONS = |- (!s t. UNIONS s INTER t = UNIONS {x INTER t | x IN s}) /\ (!s t. t INTER UNIONS s = UNIONS {t INTER x | x IN s}) Wed 13th Dec 06 Examples/hol88.ml, Examples/analysis.ml Fixed bugs in the local MK_ABS and MK_EXISTS (which is not the usual version). In both cases they had SUB_CONV instead of BINOP_CONV; this was formerly masked because until I recoded it, SUB_CONV would quietly return a reflexive theorem if the core conversion failed. Thu 7th Dec 06 term.ml Again inspired by work on Holst, recoded "vsubst" to avoid using set operations, via a generalization of "variant" (and consequently "variants" too) to avoid all free variables of the terms in the avoid list. This is doubly nicer since it avoids set operations and corresponds better to what I've actually verified in Model; those could be brought closer in step. Also, orthogonally really, made Boultonization work properly in the abstraction case. Wed 6th Dec 06 Examples/sos.ml Recoded REAL_NONLINEAR_PROVER to do a better job with trivial hypotheses. This was stimulated by a failure Laurent hit experimenting with the Coq port: REAL_SOS `&0 <= (r - r1) * (r1 - r) ==> &0 <= (r0 - r2) * (r2 - r0) ==> r1 * (r0 - r2) + r * (r2 - r2) + r1 * (r2 - r0) = &0 ==> r1 - r = &0`;; The fact that one hypothesis is trivial caused all the parameters of the ideal cofactor to "disappear" and give lookup failures afterwards. Tue 28th Nov 06 bool.ml Recoded GEN in a more elegant and efficient way. The old version was an MP relic, and it didn't do alpha-conversion very cleverly. Also tweaked EXISTS to use PROVE_HYP rather than MP. Nearly did the same to CHOOSE, but I was getting too tired to figure out for sure whether we can guarantee that PROVE_HYP will be non-trivial (all this is on the train back to Nice). Did do the same to DISJ1 and DISJ2; also fixed the exception from the latter, which was a cut-and-paste error of "DISJ1", and a similar one in "NOT_INTRO" which was a cut-and-paste "NOT_ELIM". Tue 28th Nov 06 equal.ml Recoded SUB_CONV, BINOP_CONV and DEPTH_BINOP_CONV to be a little more delicate, not that it really matters. Mon 20th Nov 06 define.ml Added clauses to scrub distinctness theorems of the form "2 * n + 1 = 2 * m + 1". This was a bit ad hoc to make "recounting the rationals" work as given, but it's reasonable as a general idea. Sun 19th Nov 06 thm.ml Recoded several primitive rules to use pattern-matching rather than the destructor functions. The efficiency gain seems real though unspectacular (around 4%) and in any case the result is slightly shorter and arguably clearer. Sun 19th Nov 06 thm.ml Changed "rator" and "rand" to use pattern-matching. Also noticed this as a result of Holst. Fri 17th Nov 06 term.ml Slightly recoded variant to use pattern-matching and the Var constructor directly (might as well since we're here!) Came across these while starting Holst on the train from Nice to Lyon. Mon 23rd Oct 06 iter.ml Fixed some faulty trivial quantifiers in NUMSEG_COMBINE_L and NUMSEG_COMBINE_R, pointed out by Lars Schewe. Mon 9th Oct 06 lib.ml, nets.ml Decommissioned the orderings like "=?" and " &0 < x / y`;; REAL_SOSFIELD `&0 <= x /\ &0 < y ==> &0 <= x / y`;; REAL_SOSFIELD `&1 < y /\ ~(z = &0) /\ abs z <= x ==> ~(x <= x / y)`;; I also noticed that there's an unused expression in the existing REAL_FIELD (I must have condensed it into the next expression and not deleted the abbreviation) which I deleted. Thu 13th Jul 06 bool.ml, calc_int.ml, calc_rat.ml, sys.ml, hol.ml, Examples/poly.ml Fixed the permissions on these files, which were unreadable for "other" users. This was pointed out by Nobuki Takayama as part of the ICMS preparation. ********************** RELEASE OF VERSION 2.20 ********************** Wed 17th May 06 tactics.ml After doing enough regression testing (particularly no changes at all to the Jordan proof), I plucked up the courage and inserted the new CONJUNCTS_THEN2 implementation which ASSUMEs the two conjuncts. I hope this will be a less invasive way of counteracting the counterintuitive effects of spurious free variables "hiding" in the goal. Wed 17th May 06 class.ml Added a quite distinct benignity mechanism and a list "the_specifications" (the old one was lost in the last change and perhaps wasn't entirely satisfactory since the original plan was not to expose the underlying definition). Wed 17th May 06 drule.ml, pair.ml Moved all the benignity, including the definition of the initial (now somewhat larger) list "the_definitions", into "pair.ml" so it directly returns the theorems from that, which is what's mainly used later. Tue 16th May 06 drule.ml Changed "the_definitions" to be as entered by the user, since Steve Brackin preferred this and I don't have any reason for this other than checking benignity. Tue 16th May 06 realax.ml Changed one proof not to rely on the fact that CONJUNCTS_THEN doesn't use ASSUMEd conjuncts. This is in the hope that I may actually manage to switch over to that, avoiding most of the confusion currently caused. Fri 12th May 06 parser.ml Changed the internal definition of "pfrees" used in set abstractions to check also whether some name has an interface mapping and if so count it as a constant. Otherwise there could be issues with overloaded constants being treated as variables in set abstractions, e.g. `onorm (f:real^M->real^N) = sup { norm(f x) | norm(x) = &1 }`;; Fri 12th May 06 preterm.ml Added a flag "type_invention_warning" to control the "inventing type variables" warning. This was inspired by Cezary, who wanted to disable these warnings and had rewritten some of the code himself to remove it; this would now have required even more duplicate-and-tweak now that "stvs_translated" is hidden. Thu 11th May 06 realax.ml Removed "real_lift_function", replacing the only uses by mapping the definition over lists. Thu 11th May 06 define.ml Renamed "pure_prove_general_recursive_function_exists" as simply "pure_prove_recursive_function_exists". The name was too long for the reference manual banner. Of course that's not such a good reason... Thu 11th May 06 tactics.ml Hid "print_hyp" and "print_hyps" inside "print_goal" and "print_goalstate" inside "print_goalstack". Thu 11th May 06 realax.ml Also removed "real_lift_theorem", replacing the only use by its definition. Thu 11th May 06 drule.ml Hid "match_bvs" inside "PART_MATCH"; it was not used anywhere else. And hid "tryalpha" inside "deep_alpha". And deleted "unify_terms". I'm a bit mystified why that's there as well as "term_unify"; the latter is at least used! Wed 10th May 06 simp.ml, Examples/mizar.ml Removed LIMITED_REWRITE_CONV, which was an egregious hack, and put in in Examples/mizar.ml, which contains the only use of it. Wed 10th May 06 thm.ml, meson.ml Removed "le_thm" and replaced the one application with its definition. Wed 10th May 06 realax.ml Hid "hreal_lift_fn" and "hreal_lift_thm". Wed 10th May 06 simp.ml Hid GEN_SUB_CONV, IMP_REWRITES_CONV and RUN_SUB_CONV. Wed 10th May 06 parser.ml Added a "string" type constraint into "install_parser", which otherwise has an undetermined type for the tags. Wed 10th May 06 ind-types.ml Removed RECTYPE_EQ_CONV, which was never used in the latest "define". Wed 10th May 06 bool.ml, passim Added "mk_iff" and "dest_iff" and renamed "is_beq" to "is_iff". Pity to make an incompatible change but it seems far more intuitive. Wed 10th May 06 ind-types.ml Again hid a lot of internal stuff: sucivate, SCRUB_EQUATION, justify_inductive_type_model, prove_model_inhabitation, define_inductive_type, define_inductive_type_constructor, instantiate_induction_theorem, pullback_induction_clause, finish_induction_conclusion, derive_induction_theorem, create_recursive_functions, create_recursion_iso_constructor, derive_recursion_theorem, generalize_recursion_theorem, define_type_mutual, TRIV_ANTE_RULE, ISO_EXPAND_CONV, lift_type_bijections, DE_EXISTENTIALIZE_RULE, grab_type, clause_corresponds, prove_inductive_types_isomorphic, SCRUB_ASSUMPTION, define_type_basecase, SIMPLE_BETA_RULE, ISO_USAGE_RULE, SIMPLE_ISO_EXPAND_RULE, REWRITE_FUN_EQ_RULE, define_type_nested. What used to be called "define_type_nested" is now called "define_type_raw" (which was formerly a limited and not specially useful version used in bootstrapping). Wed 10th May 06 preterm.ml, parser.ml Hid the following "internal" things: new_type_var, reset_type_num [actually deleted that completely], pretype_subst, pretype_instance, get_generic_type, istrivial, unify, typify, resolve_interface, solve, solve_preterm and stvs_translated. Removed pmk_eq, pmk_conj and split_ppair. Hid pmk_let, pmk_set_enum, pfrees, pmk_setabs, pmk_setcompr, pmk_vbinder, pmk_binder, pmk_cv, pmk_numeral, pgenvar, split_ppair, pmk_conj, pmk_exists. Added a few of these needed back to Examples/mizar.ml. Wed 10th May 06 lib.ml, preterm.ml Moved "num_of_string" from "preterm.ml" to "lib.ml". It seemed an unnatural fit, even though its primary purpose is indeed for parsing etc. Tue 9th May 06 parser.ml Hid "mk_precedence" and "parse_typed_apreterm", which seemed a bit too ad hoc. Tue 9th May 06 meson.ml Removed two less useful flags: "precheck" and "cache"; in neither case is it really plausible that they'll be changed. Renamed others with too-short names: depth, prefine, dcutin, skew and brand to meson_depth, meson_prefine, meson_dcutin, meson_skew and meson_brand. Hid a whole bunch of internal things in MESON: offinc, inferences, qpartition, reset_vars, fol_of_var, hol_of_var, reset_consts, fol_of_const, hol_of_const, fol_of_term, fol_of_atom, fol_of_literal, fol_of_form, hol_of_term, hol_of_atom, hol_of_literal, fol_free_in, fol_frees, fol_subst, fol_substl, fol_inst, fol_subst_bump, fol_inst_bump, istriv, fol_unify, fol_eq, cacheconts, checkan, insertan, fol_subst_partial, separate_insts, meson_single_expand, meson_expand_cont, meson_expand, expand_goal, solve_goal, fol_of_hol_clauses, optimize_rules, clear_contrapos_cache, make_hol_contrapos, meson_to_hol, create_equality_axioms, perform_brand_modification, POLY_ASSUME_TAC, SIMPLE_MESON_REFUTE, PURE_MESON_TAC, QUANT_BOOL_CONV, SPLIT_TAC. Mon 8th May 06 define.ml Removed "closed_prove_general_recursive_function_exists", which seemed a bit overspecialized and confusing. Hid various functions inside prove_general_recursive_function_exists, namely: prove_depth_measure_exists, INDUCTIVE_MEASURE_THEN, CONSTANT_MEASURE_THEN, GUESS_MEASURE_THEN, GUESS_WF_THEN, and GUESS_ORDERING_TAC. Hid "EXPAND_PAIRED_ALL_CONV", "SIMPLIFY_CASE_DISTINCTNESS_CLAUSES" and "FORALL_PAIR_CONV" inside "instantiate_casewise_recursion". Hid "tuple_function_existence", "is_applicative", "LAMBDA_PAIR_CONV" and "break_down_admissibility" inside "pure_prove_general_recursive_function_exists". Hid "close_definition_clauses" inside "define". Mon 8th May 06 define.ml, pair.ml Moved GABS_CONV back to "pair.ml", since it seems relatively natural and generally useful; I'm planning to hide most other internal functions in the "define" stuff. Also made it work on standard abstractions too, otherwise it's a bit unnatural as a general building-block. Fri 5th May 06 Proofrecording/diffs/*.ml Carefully went through making sure I'd changed the "copies" inside the Proofrecording library to keep them perfectly in step (modulo only the necessary changes) with the current core. Actually, every single file had at least cosmetic changes required. Thu 4th May 06 equal.ml Inspired by Steve Brackin's question on info-hol, generalized PAT_CONV so you can have multiple position-identifying variables. Thu 4th May 06 class.ml Changed "subst" to "vsubst" inside SELECT_CONV, which is trivial but will save a tiny bit of time. Thu 4th May 06 canon.ml, class.ml Moved REFUTE_THEN into the "class.ml" file, where it seems a more natural fit, and SPLIT_THEN into "meson.ml", since it's only used there and I may eventually rethink all this splitting stuff anyway. Hid SELECT_ELIM_CONV and SELECT_ELIM_ICONV inside SELECT_ELIM_TAC. Actually, I'm a bit confused about why I don't just do REDEPTH_CONV SELECT_CONV instead of this custom SELECT_ELIM_CONV, but I left it as is just in case I'm forgetting something important. Renamed CONDS_ELIM_CONV' as CONDS_CELIM_CONV, which is a bit more consistent with NNFC_CONV. Hid get_heads, get_thm_heads, GEN_FOL_CONV and FOL_CONV inside ASM_FOL_TAC. Wed 3rd May 06 canon.ml, meson.ml Added quick getout clauses for the reflexive case in CONJ_ACI_RULE and DISJ_ACI_RULE. Removed PROP_CNF_CONV and PROP_DNF_CONV. Renamed STRONG_CNF_CONV and STRONG_DNF_CONV to just CNF_CONV and DNF_CONV, and made them descend inside the two core quantifiers so that they subsume the old PROP_CNF_CONV and PROP_DNF_CONV. Also added ASSOC_CONV. Changed MESON to use WEAK_CNF_CONV and then ASSOC_CONV, since that seems more compatible and slightly faster than CNF_CONV. Tue 2nd May 06 canon.ml Modified PRENEX_CONV so it also pulls out existential quantifiers. I tend to use it after Skolemizing so there's no need, but for some sort of general entry point it seems sensible to enhance things. I don't think the lookup time will make a significant difference to anything. Tue 2nd May 06 canon.ml Removed MINISCOPE_CONV. In principle it's quite useful, but it's a bit overspecialized (only does universal quantifiers...) and was never used anywhere. Tue 2nd May 06 canon.ml Added an initial primitive miniscoping to SKOLEM_CONV, to push quantifiers in naively first before doing the Skolemizing, for which the universal quantifiers get pulled out again. I'm not 100% sure this is good / worthwhile, but let's try. Tue 2nd May 06 canon.ml Removed "CNNF_CONV" (NNF with an atom base conversion), which was never used; in the rare cases when this generality is needed, one can use GEN_NNF_CONV. Mon 1st May 06 int.ml, Examples/prime.ml, Examples/pocklington.ml Put in definitions of divisibility, coprimality and gcd over both Z and N. The net new definitions/theorems are, I believe: int_divides = |- !b a. a divides b <=> (?x. b = a * x) int_mod = |- !n x y. mod n x y <=> n divides x - y int_congruent = |- !x y n. (x == y) (mod n) <=> (?d. x - y = n * d) int_coprime = |- !a b. coprime (a,b) <=> (?x y. a * x + b * y = &1) num_divides = |- !a b. a divides b <=> &a divides &b num_mod = |- !n x y. mod n x y <=> mod &n (&x) (&y) num_congruent = |- !x y n. (x == y) (mod n) <=> (&x == &y) (mod &n) num_coprime = |- !a b. coprime (a,b) <=> coprime (&a,&b) num_gcd = |- !a b. gcd (a,b) = num_of_int (gcd (&a,&b)) NUM_GCD = |- !a b. &(gcd (a,b)) = gcd (&a,&b) INT_GCD_EXISTS = |- !a b. ?d. d divides a /\ d divides b /\ (?x y. d = a * x + b * y) INT_GCD_EXISTS_POS = |- !a b. ?d. &0 <= d /\ d divides a /\ d divides b /\ (?x y. d = a * x + b * y) int_gcd = |- !a b. &0 <= gcd (a,b) /\ gcd (a,b) divides a /\ gcd (a,b) divides b /\ (?x y. gcd (a,b) = a * x + b * y) Also defined a mapping "num_of_int" from integers back to natural numbers and proved a few lemmas about it: num_of_int = |- !x. num_of_int x = (@n. &n = x) NUM_OF_INT_OF_NUM = |- !n. num_of_int (&n) = n INT_OF_NUM_OF_INT = |- !x. &0 <= x ==> &(num_of_int x) = x NUM_OF_INT = |- !x. &0 <= x <=> &(num_of_int x) = x Plus a few auxiliary lemmas that I might just want to hide, but perhaps they'll be useful. FORALL_UNCURRY = |- !P. (!f. P f) <=> (!f. P (\a b. f (a,b))) EXISTS_UNCURRY = |- !P. (?f. P f) <=> (?f. P (\a b. f (a,b))) WF_INT_MEASURE = |- !P m. (!x. &0 <= m x) /\ (!x. (!y. m y < m x ==> P y) ==> P x) ==> (!x. P x) WF_INT_MEASURE_2 = |- !P m. (!x y. &0 <= m x y) /\ (!x y. (!x' y'. m x' y' < m x y ==> P x' y') ==> P x y) ==> (!x y. P x y) Also added INTEGER_TAC, INTEGER_RULE, NUMBER_TAC and NUMBER_RULE, which are a first "production" version of the ideal-based hacks I've been using lately. They will probably get some future refinements (especially the number ones, which could probably have their Finally, switched the definition of integers from "new_basic_type_definition" to just "new_type_definition" for benignity. I modified the existing files where such divisibility concepts (over N) were defined, so they use the new definitions and often use NUMBER_TAC instead of manual proofs, but are otherwise identical in structure. I can no doubt substantially improve this, but it's not critical to get to it soon. Mon 1st May 06 class.ml Fixed up "new_specification" to inherit acceptance of benign redefinition from the underlying "new_definition". This meant first not adding an outer check, and second replacing the variable in the input term by a constant if need be. Mon 1st May 06 arith.ml, realax.ml, int.ml Removed all the congruence-related stuff from its various points, as a prelude to a nicer arrangement afterwards. Mon 1st May 06 canon.ml Removed GEN_NNFC_CONV as a separate function, instead adding the same flag argument used internally to GEN_NNF_CONV. So GEN_NNF_CONV |-> GEN_NNF_CONV false GEN_NNFC_CONV |-> GEN_NNF_CONV true Also removed the toplevel binding of GEN_NNF_DCONV, which is a bit over-refined and I never actually used. Fri 28th Apr 06 list.ml Added ITLIST_APPEND = |- !f a l1 l2. ITLIST f (APPEND l1 l2) a = ITLIST f l1 (ITLIST f l2 a) The existing ITLIST_EXTRA is just a special case of this, so possibly I should remove it. Still, it's used once. Thu 27th Apr 06 help.ml [new file], doc-to-help.sed [new file] Introduced simple "help" system on the lines of HOL88. So far it's a little more inflexible, e.g. doesn't have a separate help path, but I might think about later generalizations. Couldn't resist adding a cool hack to compute the "edit distance" (aka Levenshtein distance) between two strings and try to guess what you meant. Thu 27th Apr 06 canon.ml Corrected DISJ_CANON_CONV, which simply didn't work, since several things were not modified from cut-and-pasting CONJ_CANON_CONV. Thu 27th Apr 06 canon.ml Changed DEPTH_CONV to TOP_DEPTH_CONV in PRESIMP_CONV, since several of the transformations are inherently top-down. Thu 27th Apr 06 preterm.ml, int.ml, realax.ml, real.ml Added a new function "prioritize_overload" that essentially prioritizes the first instance of each thing where the desired type appears as an instantiation of one of the type variables in the generic type. Made prioritize_num, prioritize_int, prioritize_real all just instances of this. The main point is that then it "automatically" expands to cover newly defined overloaded constants without continual redefinition of the prioritizer functions. Also renamed "mod_nat" as "nat_mod"; "mod_int" as "int_mod", and "mod_real" as "real_mod"; these are more consistent with the usual naming conventions. Also redefined "real_mod", which was stupidly trivial; moved its definition into "int.ml" so I could use the "is_int" property. Fri 21st Apr 06 lib.ml A relatively trivial change: modified the implementation of "allpairs" to be less hacky-combinatorial and probably slightly more efficient. Not that it really matters, of course... Fri 21st Apr 06 int.ml Dumped in all the operations on integers, analogous to the ones over reals: INT_LE_CONV INT_LT_CONV INT_GE_CONV INT_GT_CONV INT_EQ_CONV INT_NEG_CONV INT_MUL_CONV INT_ADD_CONV INT_SUB_CONV INT_POW_CONV INT_ABS_CONV INT_RED_CONV INT_REDUCE_CONV as well as instantiations of the normalizer and ring/ideal procedures: INT_POLY_CONV INT_RING int_ideal_cofactors Thu 20th Apr 06 sets.ml Added a stronger form of set induction, which can be useful: FINITE_INDUCT_DELETE = |- !P. P {} /\ (!s. FINITE s /\ ~(s = {}) ==> (?x. x IN s /\ (P (s DELETE x) ==> P s))) ==> (!s. FINITE s ==> P s) Thu 20th Apr 06 int.ml, Examples/floor.ml Changed "dest_int" to "real_of_int" and "mk_int" to "int_of_real". These are much more natural and people may really want to use them. I now removed interface mappings with similar effect from "Examples/floor.ml". The only remaining thing I should probably do is unify "integer" and "is_int", but that's a lower priority. It might also be nice to rename things like "int_le" so that the correspondence with real is perfect. Call that "int_le_def" and move the current INT_LE to int_le etc. Thu 20th Apr 06 int.ml Added "dest_intconst", "is_intconst" and "mk_intconst" as the start of a campaign to bring type "int" up to a more equal footing with "real". Wed 19th Apr 06 calc_int.ml Decomissioned "is_numconst", "dest_numconst" and "mk_numconst"; except for an odd use in the Maxima interface (now changed), these were little used. Also systematically renamed "[is|mk|dest]_intconst" as "[is|mk|dest|]_realintconst", since this is a bit more intuitive and makes room for those names for the type of integers. Mon 17th Apr 06 iter.ml Added a third congruence variant for iterated operations, this time for "iterator {x | P x}" assuming "P x" rather than "x IN {x | P x}" which is what the default will give. Mon 17th Apr 06 iter.ml Changed "[N]SUM_CMUL" to "[N]SUM_LMUL" and added right-handed variants "[N]SUM_RMUL": |- !f c s. nsum s (\x. f x * c) = nsum s f * c |- !f c s. sum s (\x. f x * c) = sum s f * c It was starting to get silly how often I was generating these by manually rewriting. Fri 14th Apr 06 grobner.ml, calc_rat.ml, Complex/complex.ml Made two significant changes to Grobner bases. The first is rather trifling; changed the existential variable in the Rabinowitsch theorems from "d" to "z". Since this puts it at or near the end of the resulting variable order, it matters! At least on one important example, it makes things significantly better, though I don't know how stable that is; I'm about to run tests. The other is more interesting: in the case where there are only positive equations (e.g. the Rabinowitsch trick has been pre-applied) and we have a true ring, I directly execute the "history" proof trace rather than create the intermediate cofactors list, which can sometimes really blow up. Fri 14th Apr 06 cart.ml Changed the ML binding of the definition of "finite_index" to "finite_index" rather than "index_def". Fri 14th Apr 06 calc_int.ml, calc_rat.ml Cleaned up a few places replacing "COMB2_CONV (RAND_CONV c) c" with just "BINOP_CONV c", which is simpler and probably slightly more efficient. Fri 14th Apr 06 drule.ml, calc_rat.ml Remove the ad-hoccery of "REAL_INT_POS_CONV" and "REAL_INT_POS_PROVE", replacing them by a potentially more generally useful function MP_CONV, which eliminates the antecedent of an implicational theorem |- p ==> q by applying a conversion, which may either return "p <=> T" or just "p". Fri 14th Apr 06 calc_rat.ml Deleted REAL_INT_RAT_UNOP_CONV, which in fact wasn't used at all. Replaced REAL_INT_RAT_BINOP_CONV by its trivial definition; actually made it more trivial by using BINOP_CONV. Also replaced REAL_RAT_INT_CONV by its definition and deleted it. Fri 14th Apr 06 sets.ml Added yet another few theorems about set bijections. I get the feeling that this disparate collection of lemmas is rather spinning out of control and needs to be rationalized. Anyway: BIJECTIONS_HAS_SIZE = |- !s t f g. (!x. x IN s ==> f x IN t /\ g (f x) = x) /\ (!y. y IN t ==> g y IN s /\ f (g y) = y) /\ s HAS_SIZE n ==> t HAS_SIZE n BIJECTIONS_HAS_SIZE_EQ = |- !s t f g. (!x. x IN s ==> f x IN t /\ g (f x) = x) /\ (!y. y IN t ==> g y IN s /\ f (g y) = y) ==> (!n. s HAS_SIZE n <=> t HAS_SIZE n) BIJECTIONS_CARD_EQ = |- !s t f g. (FINITE s \/ FINITE t) /\ (!x. x IN s ==> f x IN t /\ g (f x) = x) /\ (!y. y IN t ==> g y IN s /\ f (g y) = y) ==> CARD s = CARD t Thu 13th Apr 06 grobner.ml, calc_rat.ml Put a simple Rabinowitsch parameter into RING, and made the real version use it. Indeed, it seems that this can make quite a difference. One example I tried went from a minute (generating a degree-3 strong certificate in the biggest subcase) to 3 seconds. There may be some still more dramatic cases. Thu 13th Apr 06 calc_rat.ml, Complex/complex.ml Tweaked the implementations of REAL_FIELD and COMPLEX_FIELD so they do splitting up first, and separately introduce the inversion hypotheses for each conjunct. In some cases, this may save quite a few cases. For example, on the cubic, we went down from 68 cases to 32. This is still not a complete answer to its inefficiency, though: to cure that we need to tackle the underlying procedure. Perhaps I'll try aplying the Rabinowitsch trick first to avoid the overhead generating strong Nullstellensatz certificates? Thu 13th Apr 06 arith.ml, grobner.ml Finally got round to tracing why the ARITH_TAC calls in Zagier's 2-squares proof are so excruciatingly slow. Realized that the forms of theorems like SUB_ELIM_THM are stupid, because they have two separate instances of P on the right, "P d" and "P 0". It's much better to keep them as "P d" but have a "d = 0" assumption. Otherwise, you may be creating twice as many other instances, and getting doubly exponential performance! Accordingly, changed the elimination theorems and incorporated them into NUM_SIMPLIFY_CONV. Also made it behave better with respect to formula sense when introducing quantifiers, by also defining "existential" variants: PRE_ELIM_THM = |- P (PRE n) <=> (!m. n = SUC m \/ m = 0 /\ n = 0 ==> P m) PRE_ELIM_THM' = |- P (PRE n) <=> (?m. (n = SUC m \/ m = 0 /\ n = 0) /\ P m) SUB_ELIM_THM = |- P (a - b) <=> (!d. a = b + d \/ a < b /\ d = 0 ==> P d) SUB_ELIM_THM' = |- P (a - b) <=> (?d. (a = b + d \/ a < b /\ d = 0) /\ P d) DIVMOD_ELIM_THM = |- P (m DIV n) (m MOD n) <=> (!q r. n = 0 /\ q = 0 /\ r = 0 \/ m = q * n + r /\ r < n ==> P q r) DIVMOD_ELIM_THM' = |- P (m DIV n) (m MOD n) <=> (?q r. (n = 0 /\ q = 0 /\ r = 0 \/ m = q * n + r /\ r < n) /\ P q r) Thu 13th Apr 06 parser.ml Took the plunge and rewrote the precedence parser so it collects all operators with the same parse status in bunches. This was stimulated by the fact that after the recent modifications, something in Tom's Jordan proof failed because when using a binary operator "+." with the same precedence as "+", the input "a +. b + c" was parsed as "(a +. b) + c". Wed 12th Apr 06 parser.ml, preterm.ml Added support for type abbreviations: new_type_abbrev remove_type_abbrev type_abbrevs These are applied each time the type parser is looking at an atomic string. Wed 12th Apr 06 hol.ml Tweaked "load_on_path" so it only adds to the list of loaded files after the load. In some ways this seems more reasonable, since "already loaded" is a bit misleading if the load failed. Of course, one advantage of doing things the old way would be the impossibility of infinite loops... Wed 12th Apr 06 Examples/pocklington.ml Added a completely stupid factoring algorithm as a catchall; this is invoked either for small numbers (< 2^25 currently; the naive algorithm starts to take of the order of a second for primes of this size) or when the call to PARI/GP fails, e.g. because it's not installed. Tue 11th Apr 06 bool.ml Took away the initial infix of "=", which gets overwritten anyway. Very trivial but... Tue 11th Apr 06 iter.ml Added some theorems that are a bit specialized but capture quite a common situation, where you iterate something composed with an injective function from a finite set into itself: ITERATE_INJECTION = |- !op. monoidal op ==> !f p s. FINITE s /\ (!x. x IN s ==> p x IN s) /\ (!x y. x IN s /\ y IN s /\ p x = p y ==> x = y) ==> iterate op s (f o p) = iterate op s f NSUM_INJECTION = |- !f p s. FINITE s /\ (!x. x IN s ==> p x IN s) /\ (!x y. x IN s /\ y IN s /\ p x = p y ==> x = y) ==> nsum s (f o p) = nsum s f SUM_INJECTION = |- !f p s. FINITE s /\ (!x. x IN s ==> p x IN s) /\ (!x y. x IN s /\ y IN s /\ p x = p y ==> x = y) ==> sum s (f o p) = sum s f Tue 11th Apr 06 preterm.ml Enhanced "num_of_string" to handle binary numbers with notation "0b". Also put in an error check for empty digit lists in the based case; previously, for example, "0x" was quietly accepted as zero! Mon 10th Apr 06 grobner.ml Fixed a strikingly trivial error in the ideal part of RING_AND_IDEAL_CONV, where I'd missed a prime from "pol'", and so was failing by seeing that the original polynomial is nonzero. Mon 10th Apr 06 cart.ml, ind-types.ml Removed the "inductive type" definitions of the types ":2" and ":3". Added a function "define_finite_type", and made ":2" and ":3" instances of that. Derived all these theorems, some for the type ":1" which is defined the same old way: HAS_SIZE_1 = |- (:1) HAS_SIZE 1 HAS_SIZE_2 = |- (:2) HAS_SIZE 2 HAS_SIZE_3 = |- (:3) HAS_SIZE 3 DIMINDEX_1 = |- dimindex (:1) = 1 DIMINDEX_2 = |- dimindex (:2) = 2 DIMINDEX_3 = |- dimindex (:3) = 3 Mon 10th Apr 06 cart.ml Added DIMINDEX_UNIQUE = |- (:A) HAS_SIZE n ==> dimindex (:A) = n Mon 10th Apr 06 theorems.ml, passim. Finally added EQ_IMP = |- (a <=> b) ==> a ==> b which I seem to use a great deal as a prelude to congruence reasoning. Changed numerous explicit TAUT instances into that. Mon 10th Apr 06 arith.ml Added two occasionally useful and very natural rewrites: EVEN_SUB = |- !m n. EVEN (m - n) <=> m <= n \/ (EVEN m <=> EVEN n) ODD_SUB = |- !m n. ODD (m - n) <=> n < m /\ ~(ODD m <=> ODD n) Mon 10th Apr 06 iter.ml Added some theorems about iterations over sets with deleted elements. Note that the REAL one is a bit different, presumably more convenient in a ring. ITERATE_DELETE = |- !op. monoidal op ==> !f s a. FINITE s /\ a IN s ==> op (f a) (iterate op (s DELETE a) f) = iterate op s f NSUM_DELETE = |- !f s a. FINITE s /\ a IN s ==> f a + nsum (s DELETE a) f = nsum s f SUM_DELETE = |- !f s a. FINITE s /\ a IN s ==> sum (s DELETE a) f = sum s f - f a ITERATE_DELETE NSUM_DELETE Mon 10th Apr 06 arith.ml, grobner.ml Hid NUM_MULTIPLY_CONV inside NUM_SIMPLIFY_CONV; that's actually the only place where it's used. Mon 10th Apr 06 grobner.ml Completely hid almost all the internal functions inside RING_AND_IDEAL_CONV. (After carefully confirming that they weren't used anywhere, or anywhere interesting anyway.) So now all of the following are hidden away: morder_lt morder_le morder_gt grob_neg grob_add grob_sub grob_mmul grob_cmul grob_mul grob_inv grob_div grob_pow mdiv mlcm reduce1 reduceb reduce orthogonal spoly monic forder poly_lt align poly_eq memx criterion2 constant_poly grobner_basis grobner_interreduce grobner grobner_refute resolve_proof grobner_weak grobner_ideal grobner_strong Sat 8th Apr 06 grobner.ml Fixed a little error in "grobvars", which was descending to the "x" in all terms of the form "x pow n", without checking if "n" is a numeral. This was inconsistent with the actual normalizer functions, which treat "x pow n" for non-numeral n as atomic variables. For example this now works, whereas before: # REAL_RING `!x:real. &2 pow n = x ==> x = &2 pow n`;; Exception: Failure "grobify_term: unknown or invalid term". Fri 7th Apr 06 printer.ml Added ".." and "$" to the unspaced_binops. Don't know why I didn't think of it before. Fri 7th Apr 06 arith.ml, sets.ml, Examples/analysis.ml Fixed more cases of "=" that should be "<=>"; the error was only apparent because of the slight reshuffling of the precedences. In fact I discovered all but one by deliberately upping the precedence of "=" to 13 temporarily. Fri 7th Apr 06 parser.ml Modified the sorting of the infixes list to make it a canonical ordering, lexicographically by precedence, then fixity (left is higher) then alphabetical name of the operator. This was motivated by the fact that adding "CROSS" the other day had actually changed the order of "<" and "=" and affected parsing, albeit only on formulas that were "wrong". Fri 7th Apr 06 normalizer.ml Removed SEMIRING_NORMALIZE_CONV; after all it's only a small wrap round SEMIRING_NORMALIZERS_CONV, and was only used once in the core. Thu 6th Apr 06 Examples/combin.ml [new file] Added a new example, the old combinatory logic one done by Tom Melham and Juanito Camilleri in HOL88. My source and inspiration was the HOL4 tutorial and distribution; the proof is quite close to the one there apart from cosmetic features. Thu 6th Apr 06 printer.ml Restored behaviour of "name_of" that I'd changed inadvertently in my recoding: it should return "" rather than fail on combinations or abstractions. Thu 6th Apr 06 sets.ml Added a proper "Cartesian product" operation CROSS and the theorems: CROSS = |- !s t. s CROSS t = {x,y | x IN s /\ y IN t} IN_CROSS = |- !x y s t. x,y IN s CROSS t <=> x IN s /\ y IN t HAS_SIZE_CROSS = |- !s t m n. s HAS_SIZE m /\ t HAS_SIZE n ==> s CROSS t HAS_SIZE m * n FINITE_CROSS = |- !s t. FINITE s /\ FINITE t ==> FINITE (s CROSS t) CARD_CROSS = |- !s t. FINITE s /\ FINITE t ==> CARD(s CROSS t) = CARD s * CARD t All these are just trivial rewrites of "..._PRODUCT" theorems. I would rather like to decomission those, but it might be a bit disruptive. One idea would be to rename the "DEPENDENT" versions, since in many "crude" applications they would work equally well. But then all the explicit enumerations eliminated with IN_ELIM_PAIR_THM would need to be tracked. Scarcely worth it. Thu 6th Apr 06 recursion.ml, pair.ml Hid "prove_raw_recursive_functions_exist" and "prove_canon_recursive_functions_exist" inside the main function; they are never used and there's no obvious reason why one would want to. Also hid "projection_cache" and "create_projections", moving them inside GEN_BETA_CONV. Again, this seems sensible since they're hardly of general utility. I hesitated briefly because there's now no way to flush the cache, but since in the typical case it will contain only the theorem for pairs, and anyway is sorta linear in the number of constructors declared, I decided this was a non-issue. Wed 5th Apr 06 realarith.ml, calc_rat.ml Rearranged GEN_REAL_ARITH to be further specialized twice from its bootstrapping version so that it only takes the prover as an argument and otherwise has everything fixed. That makes things a bit simpler conceptually, and moreover I could now hide ABSMAXMIN_ELIM_CONV1 and ABSMAXMIN_ELIM_CONV2 inside and intermediate version. Tue 4th Apr 06 printer.ml Deleted "backquote_char" (never used, a relic from the separate filter for CAML Light), "IS_BINDER", "IS_PREFIX", "IS_INFIX", "FIXITY" (all have trivial synonyms and are only used in the printer). Hid "DEST_BINARY", "ARIGHT" and "dest_binder_vorc" inside "pp_print_term", which is the only place they are used. Renamed "NAME_OF" to "name_of" since the case conventions make the former look like an inference rule; also slightly optimized the implementation. Hid "reverse_interface". The parametrization by the reference flag makes it less generally useful, and as a matter of fact it isn't used anywhere. Also hid the initial "string_of_type" and added a new variant in the same style as the other "to string" stuff. Eventually I would like this to be derived, if at all, from the printer rather than vice versa, so it works more nicely on huge types. Sun 2nd Apr 06 int.ml Removed ARITH_CONV and INT_ARITH_CONV; these were almost never used and seem merely likely to create more confusion given that ARITH_RULE and INT_ARITH both exist and do almost the same thing. Sat 1st Apr 06 sets.ml, define.ml Got rid of various explicit "CONV_TAC(ONCE_DEPTH_CONV GEN_BETA_CONV)" type things now that generalized beta conversion is in the basic rewrites. Fri 31st Mar 06 pair.ml, realax.ml, wf.ml Removed the horribly ad-hoc GEN_PAIR_TAC. Replaced the instances in other files with other proofs, actually often simpler, typically using FORALL_PAIR_THM. Also forced "LET_TAC" to have abbreviated type "tactic". Tue 28th Mar 06 realax.ml Hid "DIST_ELIM_TAC" inside some proofs. The net effect is also a bit shorter, since I noticed three proofs were identical. Tue 28th Mar 06 sets.ml Renamed yesterday's theorem (CARD_EQ_BIJ -> CARD_EQ_BIJECTION), and added a new version with two bijections: CARD_EQ_BIJECTION = |- !s t. FINITE s /\ FINITE t /\ CARD s = CARD t ==> ?f. (!x. x IN s ==> f x IN t) /\ (!y. y IN t ==> (?x. x IN s /\ f x = y)) /\ (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) CARD_EQ_BIJECTIONS = |- !s t. FINITE s /\ FINITE t /\ CARD s = CARD t ==> ?f g. (!x. x IN s ==> f x IN t /\ g (f x) = x) /\ (!y. y IN t ==> g y IN s /\ f (g y) = y) Tue 28th Mar 06 wf.ml Got rid of WF_INDUCT_THEN, which was never used. Kept WF_INDUCT_TAC. Tue 28th Mar 06 arith.ml Got rid of PRE_ELIM_TAC and SUB_ELIM_TAC too. They're both subsumed really by NUM_MULTIPLY_CONV, and I only had one instance of each (one in Tom's Jordan proof, one in my ancient Ramsey proof where it was doing manually something doable by ARITH_TAC anyway). Tue 28th Mar 06 calc_num.ml Removed one instance of NUM_CANCEL_CONV. I had wanted to get rid of it, but then some of the proofs in "realax.ml" would become somewhat tedious, so I thought better of it. Tue 28th Mar 06 Rqe/make.ml Deleted the files "core.ml" and "analysis.ml" from the build. They aren't necessary and break when any theorems vanish, even obscure ones. Sean himself suggested getting rid of this. Mon 27th Mar 06 Examples/binary.ml [new file] Added a basic development of the "binary expansion" bijection between finite sets and functions. It's actually surprisingly tedious... Mon 27th Mar 06 arith.ml Recoded mk_small_numeral and dest_small_numeral in terms of their mk_numeral and dest_numeral counterparts. This is not only simpler but in the latter case has better error checking instead of silently overflowing. Also deleted "is_small_numeral", which is rather obscure and had never been used at all. Also hid PRE_ELIM_CONV and SUB_ELIM_CONV inside the corresponding tactics. Mon 27th Mar 06 sets.ml Removed SETIFY_CONV and SETENUM_UNION_CONV. I'm not really very satisfied with this half-baked solution. I think I should really do it all modulo an equality conversion or something. Mon 27th Mar 06 sets.ml Added another of those "how did I manage without it for so long?" theorems: CARD_EQ_BIJ = |- !s t. FINITE s /\ FINITE t /\ CARD s = CARD t ==> (?f. (!y. y IN t ==> (?x. x IN s /\ f x = y)) /\ (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y)) Mon 27th Mar 06 calc_num.ml Decided to hide NUM_DIVMOD_CONV. It's potentially useful, but I'd never used it and in any really critical situation one would probably want even more delicacy. Sun 26th Mar 06 calc_num.ml Removed toplevel binding of NUM_SUC_CONV', which is never used. Changed one right-hand clause of ARITH_LE to "n <= _0" rather than "n = 0" since it keeps the rewrites more thematic. Slightly reorganized the definitions of the inequality conversions to be more efficient exploiting this, and recoded NUM_REL_CONV. Made the final net use the explicit conversions rather than NUM_REL_CONV; I can't remember how that happened. And in the process, hid NUM_REL_CONV'. And hid both NUM_ADD_CONV' and NUM_MULT_CONV'. Finally renamed "mangle" to a more meaningful name DENUMERAL. I could just hide it, since it's not used outside this file, but that's a bit tedious. Fri 24th Mar 06 class.ml, Examples/hol88.ml Tweaked ETA_CONV so it doesn't muss the variable name. Previously ETA_CONV `\n. SUC n` returned `(\x. SUC x) = SUC`. Also removed EXT, or moved it to Examples/hol88.ml; I've never used it. And removed "list_mk_select" and "strip_select", which were never used, and "simple_new_specification". Fri 24th Mar 06 sets.ml, cart.ml [new file] Slightly reorganized the definitions of finite_index and finite_sum to use simple correspondence with ranges, parametrized by the base type, rather than the somewhat laborious constructions within the core types. As a consequence, I needed to move this stuff after the definition of ranges "..", so I put it in a new file "cart.ml". This is arguably a structural improvement anyway since it really is rather out of character with the rest of "sets.ml". Fri 24th Mar 06 ind-defs.ml Removed HALF_BETA_EXPAND; it hardly seemed worthwhile given that the definition in terms of RIGHT_BETAS is simple and it's not used much. Fri 24th Mar 06 ind-defs.ml, ind-types.ml, num.ml Deleted "prove_nonschematic_inductive_relations_exist", and changed the only two applications of it into simple "prove_inductive_relations_exist". Also removed RULE_INDUCT_TAC and RULE_INDUCT_THEN. I don't really want these ad-hoc things when MATCH_MP_TAC works fine. There are some instances of RULE_INDUCT_TAC in Examples/rstc.ml but that's already using a slightly different local definition. Also hid "generalize_schematic_variables", "derive_existence", "make_definitions", "prove_inductive_properties" and "unschematize_clauses". Fri 24th Mar 06 ind-defs.ml, class.ml, list.ml Restructured the monotonicity-proving code somewhat, hiding BACKCHAIN_TAC, MONO_ABS_TAC, APPLY_MONOTAC and MONO_STEP_TAC, and replacing "mono_tactics" with "monotonicity_theorems", which is *just* a list of theorems. This seems to be a much more straightforward interface. Thu 23rd Mar 06 ind-defs.ml Hid: FORALL_IMPS_CONV, AND_IMPS_CONV, SIMPLE_DISJ_PAIR, canonicalize_clause, canonicalize_clauses, derive_canon_inductive_relations; all of these are just used once inside the core function. Wed 22nd Mar 06 ind-defs.ml Hid "getconcl". I'll do some more tomorrow... Wed 22nd Mar 06 tactics.ml Removed the binding of REFINEMENT_PROOF, which seems entirely useless. Wed 22nd Mar 06 tactics.ml Removed the type constructor "Goalstack" everywhere and just made "goalstack" a direct abbreciation of "goalstate list". It seemed just to get in the way and make things less intuitive. Tue 21st Mar 06 tactics.ml Deleted DISJ_CASES_THENL, which is quite rarely used and has a trivial definition; it's also a bit inconsistent to have this but not the corresponding thing for conjunction. Tue 21st Mar 06 tactics.ml Fixed a small but real error in CONV_TAC, which was checking for equations "x = T" without taking into account that the goal might itself be of the form "something = T" and the tactic may be returning the goal itself. Sat 18th Mar 06 drule.ml Deleted "find_matching_subterm", which was never used at all. It was probably used in a previous version of HIGHER_REWRITE_CONV. Actually, I'm quite tempted to get rid of that too: in particular its treatment of freeness in conditionals is too conservative (only the condition arm needs to be free). Fri 17th Mar 06 drule.ml Changed BETAS_CONV so that it just infers the number of reductions to make from the form of the term, rather than taking it as a separate parameter. There's really not much to be done besides following the form of the term, and that's really all the old function did. Thu 16th Mar 06 bool.ml Made a trivial syntactic tweak to the concrete syntax in the definition of unique existence, to remove excessive bracketing. Wed 15th Mar 06 equal.ml Removed SINGLE_DEPTH_CONV, which is very specialized, only really used to do an old-style SKOLEM_CONV. The latter is used in a few places so I copied the definition there, but eventually I really ought to scrub those too. Wed 15th Mar 06 equal.ml Hid all the "failure propagating" depth conversions (even removed COMB2_QCONV, which was only used in the instance COMB_QCONV) inside the core functions. This means COMB2_QCONV, COMB_QCONV, DEPTH_QCONV, ONCE_DEPTH_QCONV, REDEPTH_QCONV, SUB_QCONV, TOP_DEPTH_QCONV and TOP_SWEEP_QCONV are no longer bound. Tue 14th Mar 06 nets.ml Hid "label_to_store", "label_for_lookup", "follow" and "net_update" inside the main two functions "enter" and "lookup". Mon 13th Mar 06 basics.ml, equal.ml Changed the notion of "path" in find_path, follow_path and PATH_CONV to be a string rather than a list of strings. After all, the whole point is to be terse. Sat 11th Mar 06 basics.ml Removed "dest_cvar". It's not unreasonable a priori but isn't even used once, and I didn't want to document it. I'm plodding through reference manual entries in Atlanta airport. Fri 10th Mar 06 make.ml, hol.ml Moved "hol_version" into the "hol.ml" file, since there seems to be no reason to confine it to the built version. And also hid "nice_date" inside "startup_banner". Fri 10th Mar 06 sys.ml Removed the assignable variable "set_jrh_lexer", which probably gives a false impression that it's a Boolean flag. Just used the name in an identity function in order to force things. Fri 10th Mar 06 lib.ml, meson.ml Removed "uniq'" and "setify'" from the library file and inserted them internally in the only application, in MESON (something I eventually want to get rid of anyway). Also removed "assoc'" and "pair_equals" completely, since they are apparently never used. Thu 9th Mar 06 lib.ml, Examples/holby.ml Deleted less useful FPF functions "tryapply" and "tryapply" (had a few instances of the latter in Holby to expand by changing to "tryapplyd ... []". Also hid a lot of things that should really be considered internal functions to the implementation: map_list, foldl_list, foldr_list, apply_listd, undefine_list, define_list, combine_list, ldb, newbranch. Thu 9th Mar 06 basics.ml Realized I'd been stupid yesterday: the messiness in "subst" with genvars is used to rename variables, and it is necessary; thanks to that a proof failed. So restored the basic structure, but kept the superior handling of unchanged subterms. Thu 9th Mar 06 lib.ml Modified "lcm_num" to return 0 rather than fail if both inputs are zero. Wed 8th Mar 06 lib.ml Fixed "gcd" to always return nonnegative gcds. Now it's basically consistent with gcd_num in the absence of overflow. Removed "lcm" which was particularly vulnerable to overflow and never got used. Wed 8th Mar 06 lib.ml Removed "munion" and "msubtract" (the later is used in the Tang exponential proof so I put a copy there). Removed all the flag functions: flags, get_flag_value, new_flag, set_flag. I had intended to use these as in HOL88, but it seems somewhat pointless; better to just use assignable variables, which can then have different types (one idea would be to put them all together in a record for tidiness). Removed "abs" (integer absolute value) which is already in OCaml and "sgn", which seems pretty minimally useful since its definition is so short; indeed it was never used. Wed 8th Mar 06 lib.ml, nets.ml Took the functions "set_insert" and "set_merge" from the library and hid them inside the unique net function they are used in. They are never used anywhere else. Wed 8th Mar 06 lib.ml, passim And removed "upto" which is just a synonm for "--". Needed to change a lot of instances. And removed "gather", which is just a synonym (and a less efficient one) for "filter". Just one or two instances in use. And removed "do_list2" which was never, ever, used. Wed 8th Mar 06 lib.ml, basics.ml Recoded "subst" to use pointer-eq rather than Unchanged exceptions, and generally wrote a simpler implementation; I can't figure out why I formerly replaced with genvars and then replaced those; it just looks weird. And this was the only place where "qcomb" was used, so removed that too. Removed "qtry" too, which wasn't used at all! Also removed all the "lazy sum" stuff that's now not used: the type "lazysum" and the functions "lazify", "eager" and "eval". Wed 8th Mar 06 lib.ml, ind-defs.ml Removed the function "assoc2", mainly to avoid documenting it, and replaced its unique application in ind-defs: let vargs = shareout xargs flargs in let cargs = map (C assoc2 (rels,vargs) o fst) uncs in with let cargs = map (fun (r,a) -> assoc2 r (rels,vargs)) uncs Also moved "shareout" back into "lib.ml"; it seems it sorta belongs there. Thu 2nd Mar 06 printer.ml Re-fixed the problem of "--&n" printing; I'd accidentally knocked this out when improving the behaviour of "----" to take into account any interface. Thu 2nd Mar 06 calc_rat.ml Added REAL_INV_INV into the initial normalizations done in REAL_FIELD; this catches an issue where in the absence of this normalization one of the proofs in 100/stirling.ml was failing; it had worked before because of a lucky correlation in the older algorithm. Wed 1st Mar 06 calc_rat.ml, Complex/complex.ml Fixed an error in REAL_FIELD and COMPLEX_FIELD: they picked out a list of any inverse terms. Now they pick only those that occur free in the formula. With the earlier implementation this led only to pointless case splits, but now it can lead to failure, e.g. if the term involves "sum (\k. ... / &(FACT k))" it would attempt to establish ~(&(FACT k) = &0) in context, which is unlikely to work. Wed 1st Mar 06 itab.ml, bool.ml, tactics.ml, class.ml, Examples/hol88.ml Removed CONTRAPOS (but put it in Examples/hol88.ml). This was very seldom used and it can be done perfectly well by GEN_REWRITE_RULE I [GSYM CONTRAPOS_THM] at worst. Likewise moved NEG_DISCH. And also SELECT_ELIM and SELECT_INTRO: neither of these was *ever* used (by me). Thu 23rd Feb 06 calc_rat.ml, Complex/complex.ml Recoded REAL_FIELD and COMPLEX_FIELD to make a much less unintelligent case-split. Instead of all 2^n combinations of cases for each x = 0 or x * inv(x) = 1 for the n terms that are inverted, use the superior n+1 cases. This should be a lot faster on formulas with many inverses. Thu 23rd Feb 06 preterm.ml Renamed "istriv" as "istrivial", since the former is already used for a MESON function (not that there's much need for anyone to use either). Tue 21st Feb 06 real.ml, num.ml, Examples/analysis.ml, Examples/transc.ml Examples/card.ml, Examples/reduct.ml, Examples/wo.ml, ind-types.ml Spotted various "duplicate definition" type errors while accumulating a list of identifiers for the reference manual. Removed duplicate proofs of REAL_LE_REFL, REAL_LE_TOTAL and REAL_LE_ANTISYM, which are already proved in "realax.ml". Also removed the pointless rewrite with _0 = 0 of SUC_INJ, which doesn't involve the constant zero. Removed double proofs of SUB_LEFT_LESS_EQ and TC_CR, and the whole series WOSET_REFL, WOSET_TRANS, WOSET_ANTISYM, WOSET_TOTAL and WOSET_WELL, plus finally INTEGRAL_LE (though there was a slight difference in the first instance) and renamed CARD_MUL_ASSOC correctly (it had been called CARD_ADD_ASSOC in a cut-and-paste error). Removed definitions of p_tm and d_tm from "transc.ml", kl_tm from "wo.ml", and localized (stupidly) t_tm in "ind-types.ml". Tue 14th Feb 06 class.ml Changed the "new_specification" code to use "new_definition" rather than "new_basic_definition", otherwise the new constants don't get added to the definitions list. Mon 13th Feb 06 printer.ml Fixed a problem with the printer pointed out by Cezary Kaliszyk: the special avoidance of printing double-negations *not* as "----" was only working for some instances, and not vectors (indeed, not complex numbers either). So I changed the code a bit so that now it comprehensively looks if the reverse interface mapping is "--", which seems more satisfactory. Fri 13th Jan 06 preterm.ml Modified "pmk_let" so that it will allow "<=>" as well as "=" as the binding construct; remember that typechecking and overloading resolution hasn't been performed at this point. This was to fix an unanticipated issue with the change of precedence: neither "let x = y \/ z in ..." nor "let x <=> y \/ z in ..." was accepted. Now the latter is; in either case "let x = (y \/ z) in ..." was of course. Still, I wonder if I ought to do a more special-case parse of the let bindings rather than relying on treating it as a preterm directly. Of course it prints as "<=>" anyway. Let's stick with this for now. Mon 9th Jan 06 define.ml Fixed one case where there was a reliance on the default rewriting net at runtime, in ELIM_LISTOPS_CONV within instantiate_casewise_recursion. With paired beta inside, this was breaking. Also changed REWRITE_CONV to PURE_REWRITE_CONV on line 451. Although I don't think that was currently causing problems, it could in principle. There are still a few defaulted REWRITE_TACs in the later order-guessing stuff, but then that should be rewritten more fundamentally anyway so I'm not planning to tinker unless I need to. Fri 6th Jan 06 pair.ml So, added GEN_BETA_CONV to the basic rewrites at the end of "pair.ml". Seems not to kill anything in the core; I'll run some tests to see if anything breaks. Eventually I can go through and make some consequential simplifications. Fri 6th Jan 06 simp.ml I really wanted to add paired (or generalized) beta-conversion to the basic "rewrites", but the mechanism for adding a conversion is not really there. So I restructured, adding set_basic_convs extend_basic_convs basic_convs and tweaked all the setting functions so that updating either the basic theorems or conversions will "rehash" the whole term net starting from those two lists. Should be compatible... Wed 4th Jan 06 printer.ml Prompted by Steve Brackin, who wanted to convert terms to strings, I generalized several printing functions into "pp_" variants that take a formatter as an additional argument: pp_print_type pp_print_qtype pp_print_term pp_print_qterm pp_print_thm Also added more explicit string conversion functions: print_to_string string_of_term string_of_thm Mon 12th Dec 05 iter.ml Removed more redundant finiteness assumptions from some theorems. These may seem a bit too "hacky" to expose, but it's very convenient to be able to apply these with no finiteness worries. NSUM_SUPPORT = |- !f s. nsum (support (+) f s) f = nsum s f SUM_SUPPORT = |- !f s. sum (support (+) f s) f = sum s f NSUM_CMUL = |- !f c s. nsum s (\x. c * f x) = c * nsum s f SUM_CMUL = |- !f c s. sum s (\x. c * f x) = c * sum s f SUM_NEG = |- !f s. sum s (\x. --f x) = --sum s f Deleted NSUM_CMUL_NUMSEG, SUM_CMUL_NUMSEG and SUM_NEG_NUMSEG since they're now nothing but instances of the general case. Fri 9th Dec 05 iter.ml Realized that several "nsum" theorems had stupid "0 <= ..." properties that I just took over from the real numbers, but of course these are redundant for N. Just removed NSUM_POS_LE and NSUM_POS_LE_NUMSEG, NSUM_POS_EQ_0 and NSUM_POS_EQ_0_NUMSEG. On the other hand added the more useful bi-implications: NSUM_EQ_0_IFF = |- !s. FINITE s ==> (nsum s f = 0 <=> !x. x IN s ==> f x = 0) NSUM_EQ_0_IFF_NUMSEG = |- !f m n. nsum (m .. n) f = 0 <=> (!i. m <= i /\ i <= n ==> f i = 0) Thu 8th Dec 05 iter.ml Made quite a few theorems stronger by removing unnecessary finiteness assumptions. The first was just stupidity; the others have become possible thanks to the canonical choice in the "infinite support" case: ITERATE_SUPPORT = |- !op f s. iterate op (support op f s) f = iterate op s f ITERATE_IMAGE = |- !op. monoidal op ==> (!f g s. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> iterate op (IMAGE f s) g = iterate op s (g o f)) ITERATE_BIJECTION = |- !op. monoidal op ==> (!f p s. (!x. x IN s ==> p x IN s) /\ (!y. y IN s ==> (?!x. x IN s /\ p x = y)) ==> iterate op s f = iterate op s (f o p)) ITERATE_EQ = |- !op. monoidal op ==> (!f g s. (!x. x IN s ==> f x = g x) ==> iterate op s f = iterate op s g) ITERATE_EQ_GENERAL = |- !op. monoidal op ==> (!s t f g h. (!y. y IN t ==> (?!x. x IN s /\ h x = y)) /\ (!x. x IN s ==> h x IN t /\ g (h x) = f x) ==> iterate op s f = iterate op t g) ITERATE_EQ_GENERAL_INVERSES = |- !op. monoidal op ==> (!s t f g h k. (!y. y IN t ==> k y IN s /\ h (k y) = y) /\ (!x. x IN s ==> h x IN t /\ k (h x) = x /\ g (h x) = f x) ==> iterate op s f = iterate op t g) and likewise for all the instantiations. Also added the following clause to SUPPORT_CLAUSES: |- ... /\ (!f g s. support op g (IMAGE f s) = IMAGE f (support op (g o f) s)) Thu 8th Dec 05 sets.ml Added FINITE_IMAGE_INJ_EQ |- !f s. (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y) ==> (FINITE (IMAGE f s) <=> FINITE s) Wed 7th Dec 05 iter.ml Changed the definition of "iter" so it always returns the neutral element in the case of non-finite support. The principal motivation is to allow us to have a simple congruence rule, which at the moment we have for the special case of "m..n" but not for general sets. Another pleasing consequence is that quite a few theorems (ITERATE_EQ, ITERATE_IMAGE, ...) can lose their finiteness hypothesis. However, I haven't at present added the new congruence rules or modified the theorems; that will require quite a bit of effort to fix proofs broken by becoming "too easy". Tue 6th Dec 05 canon.ml, meson.ml Attempted to improve CONDS_ELIM_CONV, a slightly dangerous undertaking. Now it more aggressively splits at the outer level, even if the term contains quantifiers, provided the condition being tested is free. Moreover, it will achieve some level of sharing when multiple conditional expressions have the same test. I also changed CONDS_ELIM_CONV to CONDS_ELIM_CONV' in MESON, since it seems to be a case where conjunctive splitting is called for. Tue 6th Dec 05 iter.ml Added variants of general equality between sums using mutually inverse functions. This clutter is getting a bit much, but these often seem more useful. ITERATE_EQ_GENERAL_INVERSES = |- !op. monoidal op ==> (!s t f g h k. FINITE s /\ (!y. y IN t ==> k y IN s /\ h (k y) = y) /\ (!x. x IN s ==> h x IN t /\ k (h x) = x /\ g (h x) = f x) ==> iterate op s f = iterate op t g) NSUM_EQ_GENERAL_INVERSES = |- !s t f g h k. FINITE s /\ (!y. y IN t ==> k y IN s /\ h (k y) = y) /\ (!x. x IN s ==> h x IN t /\ k (h x) = x /\ g (h x) = f x) ==> nsum s f = nsum t g SUM_EQ_GENERAL_INVERSES = |- !s t f g h k. FINITE s /\ (!y. y IN t ==> k y IN s /\ h (k y) = y) /\ (!x. x IN s ==> h x IN t /\ k (h x) = x /\ g (h x) = f x) ==> sum s f = sum t g Mon 5th Dec 05 Examples/pocklington.ml It was pointed out to me by Freek that there's still no proof of the multiplicativity of the totient function, so I added this and various other useful lemmas: PHI_FINITE_LEMMA = |- !P n. FINITE {m | coprime (m,n) /\ m < n} CONG_IMP_EQ = |- !x y n. x < n /\ y < n /\ (x == y) (mod n) ==> x = y CONG_DIVIDES_MODULUS = |- !x y m n. (x == y) (mod m) /\ n divides m ==> (x == y) (mod n) MOD_MULT_CONG = |- !a b x y. ~(a = 0) /\ ~(b = 0) ==> ((x MOD (a * b) == y) (mod a) <=> (x == y) (mod a)) including more forms of the Chinese remainder theorem combining existence and uniqueness: CHINESE_REMAINDER_UNIQUE = |- !a b m n. coprime (a,b) /\ ~(a = 0) /\ ~(b = 0) ==> (?!x. x < a * b /\ (x == m) (mod a) /\ (x == n) (mod b)) CHINESE_REMAINDER_COPRIME_UNIQUE = |- !a b m n. coprime (a,b) /\ ~(a = 0) /\ ~(b = 0) /\ coprime (m,a) /\ coprime (n,b) ==> (?!x. coprime (x,a * b) /\ x < a * b /\ (x == m) (mod a) /\ (x == n) (mod b)) and finally the multiplicativity property itself: PHI_MULTIPLICATIVE = |- !a b. coprime (a,b) ==> phi (a * b) = phi a * phi b Mon 5th Dec 05 sets.ml Added CARD_IMAGE_INJ_EQ, which can be handy to avoid explicitly expanding the image. |- !f s t. FINITE s /\ (!x. x IN s ==> f x IN t) /\ (!y. y IN t ==> (?!x. x IN s /\ f x = y)) ==> CARD t = CARD s Sat 3rd Dec 05 sets.ml Added a type annotation to INTERS_0 so it doesn't look so ugly with the new union-printing. Sat 3rd Dec 05 meson.ml Tried to do the dual thing for MESON, adding a new UNWIND_CLAUSAL_EQUATION to eliminate equality from assumptions that modulo clausification are like `!x. x = a ==> P(x)` and replace them with `P(a)`. Took the chance to eliminate definition of the function GSPEC, which is only really used locally and clashes with the definitional theorem for the set constant. Just doing UNWIND_CLAUSAL_EQUATION where possible *can* actually make things worse, because the "stupid" form was useful in contrapositive mode I assume. So instead of replacing the "stupid" form I add both. This is a compromise: it may still increase the search space but at least shouldn't increase the depth and so *really* be a disaster. Well, that's the theory... Sat 3rd Dec 05 tactics.ml Generalized SUBST_VAR_TAC to do constants too, which is even easier. Fri 2nd Dec 05 tactics.ml, meson.ml Added SUBST_VAR_TAC, which will perform substitution if one side of the equation is a variable not occurring in the other (or do nothing for a reflexive equation). Added it just at the end of the MESON canonicalization; I hope this will render some simple equality propagation more efficient. ********************** RELEASE OF VERSION 2.10 ********************** Wed 30th Nov 05 Examples/permutations.ml, Examples/product.ml [new files] Added these two files, one with a quite decent theory of permutations, one with a fairly spartan theory of products, which could be beefed up with a little boring work. Wed 30th Nov 05 iter.ml Added theorems about bijections of the indexing set: ITERATE_BIJECTION = |- !op. monoidal op ==> (!f p s. FINITE s /\ (!x. x IN s ==> p x IN s) /\ (!y. y IN s ==> (?!x. x IN s /\ p x = y)) ==> iterate op s f = iterate op s (f o p)) SUM_BIJECTION = |- !f p s. FINITE s /\ (!x. x IN s ==> p x IN s) /\ (!y. y IN s ==> (?!x. x IN s /\ p x = y)) ==> sum s f = sum s (f o p) NSUM_BIJECTION = |- !f p s. FINITE s /\ (!x. x IN s ==> p x IN s) /\ (!y. y IN s ==> (?!x. x IN s /\ p x = y)) ==> nsum s f = nsum s (f o p) about composing iterated operations over a Cartesian product set: ITERATE_ITERATE_PRODUCT = |- !op. monoidal op ==> (!s t x. FINITE s /\ (!i. i IN s ==> FINITE (t i)) ==> iterate op s (\i. iterate op (t i) (x i)) = iterate op {i,j | i IN s /\ j IN t i} (\(i,j). x i j)) NSUM_NSUM_PRODUCT = |- !s t x. FINITE s /\ (!i. i IN s ==> FINITE (t i)) ==> nsum s (\i. nsum (t i) (x i)) = nsum {i,j | i IN s /\ j IN t i} (\(i,j). x i j) SUM_SUM_PRODUCT = |- !s t x. FINITE s /\ (!i. i IN s ==> FINITE (t i)) ==> sum s (\i. sum (t i) (x i)) = sum {i,j | i IN s /\ j IN t i} (\(i,j). x i j) And this, which had somehow been forgotten in the general case: ITERATE_EQ = |- !op. monoidal op ==> (!f g s. FINITE s /\ (!x. x IN s ==> f x = g x) ==> iterate op s f = iterate op s g) and these more general-looking theorems in all instantiations: ITERATE_EQ_GENERAL = |- !op. monoidal op ==> (!s t f g h. FINITE s /\ (!y. y IN t ==> (?!x. x IN s /\ h x = y)) /\ (!x. x IN s ==> h x IN t /\ g (h x) = f x) ==> iterate op s f = iterate op t g) NSUM_EQ_GENERAL = |- !s t f g h. FINITE s /\ (!y. y IN t ==> (?!x. x IN s /\ h x = y)) /\ (!x. x IN s ==> h x IN t /\ g (h x) = f x) ==> nsum s f = nsum t g SUM_EQ_GENERAL = |- !s t f g h. FINITE s /\ (!y. y IN t ==> (?!x. x IN s /\ h x = y)) /\ (!x. x IN s ==> h x IN t /\ g (h x) = f x) ==> sum s f = sum t g Wed 30th Nov 05 sets.ml Added IN_ELIM_PAIR_THM = |- !P a b. (a,b) IN {(x,y) | P x y} <=> P a b Tue 22nd Nov 05 Examples/transc.ml Added some missing derivative theorems to "diff_net": sqrt, asn and acs. The DIFF_SQRT theorem wasn't even there at all: DIFF_SQRT = |- !x. &0 < x ==> (sqrt diffl inv (&2 * sqrt x)) x Fri 18th Nov 05 Examples/floor.ml Put interface maps "real_of_int" and "int_of_real" for the "dest_int" and "mk_int" functions, and proved "integer" and "is_int" are the same. Really, I would have preferred to just fix the names back in "int.ml" but was put off by the thought of breaking Tom's proofs: he uses "dest_int" at least quite extensively and may rely on its hash value. Thu 17th Nov 05 make.ml Added a version number (now set to "2.10" for next planned public release) and set it to appear in the startup banner. Thu 17th Nov 05 meson.ml Scaled back the default chattiness of MESON (unless "meson_chatty" is set), with a one-line output that's almost as informative as the old screed. Thu 17th Nov 05 parser.ml. printer.ml Introduced "string_of_type" and "print_type" (now without surrounding quotes and colon) and renamed old "print_type" (with quotes and colon) as "print_qtype". This is more consistent with the naming scheme for terms. Introduced parsing and printing of "(:ty)" as "UNIV:ty->bool". This is very helpful when using my functional approach to Cartesian products. It can be disabled by setting the new variable "typify_universal_set" to false. Wed 16th Nov 05 define.ml Added some limited support for intelligent admissibility analysis of "sum" and "nsum". Slightly lazily, I don't do the careful tupling, and just do beta-conversion afterwards, so it's only really reliable for one-argument functions that appear themselves in the summation. But this is a useful special case. It would be easy but tedious to generalize it. The new theorems themselves are, I believe, sufficiently generic, ADMISSIBLE_SUM = |- !(<<) p s h a b. (!k. admissible (<<) (\f x. a(x) <= k /\ k <= b(x) /\ p f x) s (\f x. h f x k)) ==> admissible (<<) p s (\f x. sum (a(x) .. b(x)) (h f x)) ADMISSIBLE_NSUM = |- !(<<) p s h a b. (!k. admissible (<<) (\f x. a(x) <= k /\ k <= b(x) /\ p f x) s (\f x. h f x k)) ==> admissible (<<) p s (\f x. nsum (a(x) .. b(x)) (h f x)) Tue 8th Nov 05 pa_j_3.09.ml [new file] Flemming Andersen pointed out that the latest pa_j.ml doesn't work in 3.09; there have been yet more changes! So I created a new pa_j.ml for 3.09, correctly I hope. Tue 8th Nov 05 preterm.ml, parser.ml, printer.ml Added pmk_setcompr, which generalizes pmk_setabs to be given an explicit set of bound variables (code from Sean McLaughlin). Now "pmk_setabs" just figures out the default and calls that. In the parser, added parsing for set comprehensions with explicit bound variable indication as in "{t | vs | p}". The printer will also print in this form if either (i) flag print_unambiguous_comprehensions is set, or (ii) the choices don't match the defaults (which are unchanged). Fri 4th Nov 05 printer.ml Added two new reference variables "unspaced_binops" (binary operators to be printed without spaces round them) and "prebroken_binops" (binary operators where line breaks are to be inserted before the operator for preference). This generalizes the special treatments originally given to "," and "==>" respectively, which I wanted to extend to "::" and "-->" for a Murphi spec. Also added "reverse_interface" and used it, and changed DEST_BINARY further, so that it treats as the same all constants (only constants) that have the same interface map. Once again this is good for "::" which I had overloaded to two different things. Wed 2nd Nov 05 equal.ml Added PAT_CONV, identifying positions for conversion application using a simple pattern term. I've contemplated this for a while and I know John Matthews also suggested it. Mon 31st Oct 05 hol.ml Just when I was ready for a Halloween release! I found that I was not consistently keeping a fully qualified path in the lost "loaded_files". This can mean "needs" forcing a repetition of something that was loaded using ".". So I added an explicit transformation of "." into "Sys.getcwd()" in "file_on_path". Fri 28th Oct 05 parser.ml Realized I'd been too hasty accommodating Steve B's requests. The new lexical conventions hit me almost as soon as I started loading the complex numbers and hit `&1,&2`. So I implemented a somewhat more intricate approach, where "," and ";" are set aside into their own class of "separators", and the only permissible combination is repetition within that class. This gives ",,", which Steve wanted, and ";;", which I sometimes use in programming language semantics. Fri 28th Oct 05 calc_rat.ml Finally fixed REAL_RAT_DIV_CONV so it doesn't incorrectly ignore "&n / &1". Also fixed what is a bug, though I've never seen it come up: I was using "=" instead of "=/" for the gcd equality test. Thu 27th Oct 05 lib.ml After finally tracking a performance bug to large traversals inside set operations like "union", I defined inequalities and equalities like "=?", "<=?" etc. and tried to use them consistently inside the set operations and other related stuff like "setify", finite partial functions etc. I found this by porting by old Stalmarck code because Hasan Amjad was interested, and discovered that compared with 1995 the pure (O)Caml proof search is about 50X as fast, yet sometimes proof reconstruction was only 3X. Thu 27th Oct 05 bool.ml Reorganized derivations to avoid gratuituous MP in CONTR and DISJ_CASES. There are plenty of others worth doing, no doubt. Wed 26th Oct 05 parser.ml Moved the comma from punctuation/bracket status to regular symbol. There seems no really compelling reason for its position, given that "." and ";" are now classed as regular symbols. The only slight danger is using identifiers with leading or trailing underscores next to commas, as for example with ML pattern-matching "a,_,b". This was stimulated by Steve Brackin's desire to use ",," as the record selector. Mon 24th Oct 05 calc_rat.ml Added inv(x pow n) = inv(x) pow n to the initial normalization in REAL_FIELD, which is consistent with the existing use of inv(x * y) = inv(x) * inv(y). This fixes the fact that things like this failed (and none too fast!) REAL_FIELD `&0 < x ==> &1 / x pow 2 - &1 / (x + &1) pow 2 = (&2 * x + &1) / (x * (x + &1)) pow 2`;; Mon 24th Oct 05 Makefile Added a catchall so that if there's no pa_j_XXX.ml for the version of OCaml being used, we use the 3.08 one. This was stimulated by Steve Brackin's observation that the build didn't just work on 3.08.4. Thu 20th Oct 05 Examples/analysis.ml Added a theorem that's pretty trivial, but tiresome to prove each time it comes up: SEQ_HARMONIC = |- !a. (\n. a / &n) --> &0. Thu 20th Oct 05 Examples/hol88.ml [new file] Added a "HOL88 compatibility" file. This is by no means perfect, but fixes the most important gaps and incompatibilities I've come across when porting stuff. Wed 19th Oct 05 grobner.ml Fixed a bug in "grobner_strong" by adding a special case for pol = []. As it stood, this was correctly returning the trivial certificate, but incorrectly returning the degree d = 0 instead of d = 1. This caused some later things to fail in trivial cases, e.g. REAL_RING `x:real = y ==> z:real = z`;; Mon 17th Oct 05 real.ml Added REAL_SOS_EQ_0 = |- !x y. x pow 2 + y pow 2 = &0 <=> x = &0 /\ y = &0 Fri 7th Oct 05 parser.ml Added "isalpha" (letter or prime or underscore) to the character discrimination functions, and slightly rejigged the implementation. Tue 20th Sep 05 lib.ml Tweaked "setify" and "setify'" to remove the pointless historical relic of a trap for the Unchanged exception. And it was wrong anyway: should have returned the intermediate sorted list if the exception came from "uniq". This was pointed out by Sean McLaughlin. Wed 7th Sep 05 printer.ml At Sean's suggestion, modified the store of printers to have names. So now "install_user_printer" takes a name-printer pair not just a printer, and delete_user_printer takes just a name (it wasn't really usable with just a function since they can't be compared for equality). Thu 1st Sep 05 ind-types.ml Added induction and recursion theorems for treating "bool" as just another inductive type: bool_INDUCT = |- !P. P F /\ P T ==> (!x. P x) bool_RECURSION = |- !a b. ?f. f F = a /\ f T = b and added them to the inductive type store, with knock-on effects on the "rectype net". This was needed to make a definition Damir Jamsek wanted to work accepted by "define": define `(cnt3(T,T) = 0) /\ (cnt3(F,T) = 1) /\ (cnt3(T,F) = 2)`;; Sat 20th Aug 05 tactics.ml Added a warning to the "g" goal-setting function about free variables in the goal. Tue 16th Aug 05 parser.ml At Sean's suggestion, removed the second (parser function) argument from "delete_parser", since it wasn't used anyway. Thu 11th Aug 05 sets.ml Added some theorems about general intersections: INTERS_0 = |- INTERS {} = UNIV INTERS_1 = |- INTERS {s} = s INTERS_2 = |- INTERS {s,t} = s INTER t INTERS_INSERT = |- INTERS (s INSERT u) = s INTER (INTERS u) Tue 2nd Aug 05 Examples/analysis.ml, Examples/transc.ml Added another tranche of basic but sometimes non-trivial theorems about integration, including integration over a combined interval, integration over a subinterval and integrability of continuous functions. Tue 26th Jul 05 tactics.ml, simp.ml Enhanced ABBREV_TAC so you can introduce functions with arguments directly rather than explicitly using the lambda-abstraction. Haven't correspondingly changed EXPAND_TAC as yet, and I'm not sure I will. Also moved ABBREV_TAC and EXPAND_TAC into "simp.ml" so the former can use rewriting. Mon 25th Jul 05 Examples/analysis.ml, Examples/transc.ml Added quite a few theorems to these files. Some, amazingly, had been put into the CAML Light version (gtt) but left out here. More substantially, added quite a few new lemmas about integration, which otherwise had hardly any theorems except for the FTC. ********************** RELEASE OF VERSION 2.00 ********************** Fri 22nd Jul 05 hol.ml, Examples/sos.ml At the suggestion of Tom Hales, added "tmp_path", a settable variable for a temporary path, and used that instead of explicit "/tmp" string in SOS. Fri 22nd Jul 05 sets.ml Added HAS_SIZE_PRODUCT_DEPENDENT = |- !s m t n. s HAS_SIZE m /\ (!x. x IN s ==> t(x) HAS_SIZE n) ==> {(x:A,y:B) | x IN s /\ y IN t(x)} HAS_SIZE (m * n) and made CARD_PRODUCT proof just use this as a lemma. Fri 22nd Jul 05 Examples/pocklington.ml Added more basic lemmas about congruences, including solving linear congruences: CONG_MOD_MULT = |- !x y m n. (x == y) (mod n) /\ m divides n ==> (x == y) (mod m) CONG_SOLVE = |- !a b n. coprime(a,n) ==> ?x. (a * x == b) (mod n) CONG_SOLVE_UNIQUE = |- !a b n. coprime(a,n) /\ ~(n = 0) ==> ?!x. x < n /\ (a * x == b) (mod n) CONG_UNIQUE_INVERSE_PRIME = |- !p x. prime p /\ 0 < x /\ x < p ==> ?!y. 0 < y /\ y < p /\ (x * y == 1) (mod p) Fri 22nd Jul 05 Examples/card.ml Added CANTOR_THM = |- (UNIV:A->bool) <_c (UNIV:(A->bool)->bool) to the cardinality theory. Tue 12th Jul 05 hol.ml Changed the first element of load_path to "."; formerly it used Sys.getcwd() but of course that's at build time, making it rather useless. Tue 12th Jul 05 calc_rat.ml Fixed up REAL_FIELD to (i) avoid using the non-trivial inverse adjunction for inverses of rational constants, and (ii) do a conjunctive not disjunctive split in its initial canonicalization. Tue 12th Jul 05 grobner.ml, passim Added another parameter to "RING_AND_IDEAL_CONV" (hence also "RING" and "ideal_cofactors") for the inversion ("inv") operation of the ring. Made the corresponding routine changes to the code to exploit "inv" correctly. It seemed inconsistent that, say, REAL_RING copes fine with &1 / &3 but not inv(&3)! Tue 12th Jul 05 canon.ml Added NNFC_CONV and GEN_NNFC_CONV, which force a conjunctive NNF split. This is sometimes useful if you're trying to prove something rather than refute its negation. Mon 11th Jul 05 grobner.ml Made "grobner_ideal" fail if the polynomial is not in the ideal instead of just returning an unsatisfactory list of cofactors. Obviously this is better on general grounds, and in particular it ensures that failure comes early on big applications of the IDEAL part of RING_AND_IDEAL_CONV. Also added initial interreduction of the input polynomials before forming a Grobner basis. (It would be more elegant just to put them in the spoly rather than basis list, but that gets complicated by all the criteria that are applied only when processed.) Also added depth BETA_CONV and PRESIMP_CONV to the initializations that go on, otherwise it fails to deal properly with things like REAL_RING `x = &1 ==> T`; also put a special-case check in the REFUTE code for `F`. Finally, fixed a problem where if the Strong Nullstellensatz certificate happened to have degree zero (i.e. the equations along imply 1=0) yet there are also assumed inequations to refute, the wrong thing was generated. This line: let th2 = funpow (deg-1) (IDOM_RULE o CONJ th1) th1 in was changed to this: let th2 = funpow deg (IDOM_RULE o CONJ th1) NOT_EQ_01 Sun 10th Jul 05 sets.ml Added yet more rather elementary set-theoretic results that aren't completely trivial conseqences of stuff I already have: CHOOSE_SUBSET = |- !s:A->bool. FINITE s ==> !n. n <= CARD s ==> ?t. t SUBSET s /\ t HAS_SIZE n CARD_UNION_LE = |- !s t:A->bool. FINITE s /\ FINITE t ==> CARD(s UNION t) <= CARD(s) + CARD(t) CARD_UNIONS_LE = |- !s t:A->B->bool m n. s HAS_SIZE m /\ (!x. x IN s ==> FINITE(t x) /\ CARD(t x) <= n) ==> CARD(UNIONS {t(x) | x IN s}) <= m * n Sat 9th Jul 05 grobner.ml Fixed an error where in the "no negated equations" case of Grobner, the filtered set of "things that are indeed equations" was not being used consistently, but rather at one point the original term was fed to "grobify_equations". Fri 8th Jul 05 sets.ml Added a few new set-theoretic lemmas: EMPTY_UNIONS = |- !s. UNIONS s = {} <=> (!t. t IN s ==> t = {}) HAS_SIZE_UNION = |- !s t m n. s HAS_SIZE m /\ t HAS_SIZE n /\ DISJOINT s t ==> s UNION t HAS_SIZE m + n HAS_SIZE_UNIONS = |- !s t m n. s HAS_SIZE m /\ (!x. x IN s ==> t x HAS_SIZE n) /\ (!x y. x IN s /\ y IN s /\ ~(x = y) ==> DISJOINT (t x) (t y)) ==> UNIONS {t x | x IN s} HAS_SIZE m * n Thu 7th Jul 05 Examples/prime.ml Added DIVIDES_PRIMEPOW = |- !p. prime p ==> (!d. d divides p EXP k <=> (?i. i <= k /\ d = p EXP i)) COPRIME_DIVISORS = |- !a b d e. d divides a /\ e divides b /\ coprime(a,b) ==> coprime(d,e) Thu 7th Jul 05 Makefile Added some files to the HOLSRC list in the Makefile; it was missing a few recent ones such as "iter.ml". Thu 7th Jul 05 iter.ml Added some more theorems about sums w.r.t. subsets, which I was a bit surprised I'd never proved before: NSUM_SUBSET = |- !u v f. FINITE u /\ FINITE v /\ (!x. x IN u DIFF v ==> f x = 0) ==> nsum u f <= nsum v f NSUM_SUBSET_SIMPLE = |- !u v f. FINITE v /\ u SUBSET v ==> nsum u f <= nsum v f SUM_SUBSET = |- !u v f. FINITE u /\ FINITE v /\ (!x. x IN u DIFF v ==> f x <= &0) /\ (!x. x IN v DIFF u ==> &0 <= f x) ==> sum u f <= sum v f SUM_SUBSET_SIMPLE = |- !u v f. FINITE v /\ u SUBSET v /\ (!x. x IN v DIFF u ==> &0 <= f x) ==> sum u f <= sum v f Sun 3rd Jul 05 Examples/analysis.ml Used overloading rather than overriding for the new and "pair-style" sums so it isn't so thoroughly awkward to mix the two. And added two theorems that relate old-style and new-style sums (which are a bit messy): PSUM_SUM = |- !f m n. sum (m,n) f = sum {i | m <= i /\ i < m + n} f PSUM_SUM_NUMSEG = |- !f m n. ~(m = 0 /\ n = 0) ==> sum (m,n) f = sum (m .. (m + n) - 1) f Sun 3rd Jul 05 calc_rat.ml Added a small but significant tweak to REAL_FIELD: rewrite with REAL_INV_MUL in the initial canonicalization. This often means that it only has to prove nonzeroness for individual factors, which is in general much easier. For example, it can now do this, which it couldn't before: REAL_FIELD `&1 / (&n + &1) - &1 / (&n + &2) = &1 / ((&n + &1) * (&n + &2))`;; Sun 3rd Jul 05 iter.ml Added: REAL_OF_NUM_SUM = |- !f s. FINITE s ==> &(nsum s f) = sum s (\x. &(f x)) REAL_OF_NUM_SUM_NUMSEG = |- !f m n. &(nsum (m .. n) f) = sum (m .. n) (\i. &(f i)) Fri 1st Jul 05 Examples/analysis.ml, Examples/transc.ml Changed the definitions of "sup" and "sqrt" to match the ones in the multivariate theory (called xxx_def here) and changed the current definitional theorems "sup" and "sqrt" to derived consequences. This removes some of the more gratuitous clashes between the univariate and multivariate theories. Thu 30th Jun 05 iter.ml, real.ml Changed "CARD_EQ_SUM" to "CARD_EQ_NSUM" for the "nsum" version; before it erroneously duplicated the "sum" version. Also removed duplicates from "real.ml" of REAL_LE_ADD, REAL_LET_ADD and REAL_LT_ADD, already defined in "realarith.ml". All this name duplication was found by Steven Obua's editing script! Thu 30th Jun 05 lib.ml, passim Made changes so that Steven Obua's proof recording stuff does not require much to be changed. Mainly: add new parametrized set operations in "lib.ml", define "equals_thm" and "equals_goal" and use them consistently to avoid ever comparing theorems for equality with the built-in relation. With a few minor changes, all the things I did were just following Steven's suggestions. Wed 29th Jun 05 inter.ml Changed some bound variables names that range over number segments from x/y (as in the general set case) to i/j, which seems a bit more intuitive. Tue 28th Jun 05 define.ml Fixed some degenerate cases in "define" and associated functions so that it can handle a trivial equational definition "v = t" and hence be used quite generally. Tue 28th Jun 05 tactics.ml Finally removed the use of "mk_thm" from tactic validity checking. Instead used a technique going back to Cambridge LCF: added a function "mk_fthm" which adds a trivially false assumption (actually a distinct constant _FALSITY_ to avoid unfortunate interactions with real parts of the goal). On a superficial test it seems to work fine. Mon 27th Jun 05 class.ml, ind-types.ml Added sort of benignity checking to type definitions made with "new_type_definition" and inductive types defined with "define_type". They are separate mechanisms, and the second is rather crude, actually using the input *string*. But it should suffice to make sure that almost all proof files will now load multiple times without redefinition complaints. Mon 27th Jun 05 drule.ml, define.ml Added "benign definition" acceptance to "define" and beefed up that in "new_definition" with the same new feature: it's OK it type variables etc. change provided both the old and new definitions can be instantiated (by PART_MATCH) to each other. (This ensures that we can change type variables but not make the definition more type-constrained, e.g. making two formerly distinct type variables the same.) Mon 27th Jun 05 tactics.ml Changed "REMOVE_THEN" so it removes only the first assumption with that label (and always the one that's handed to the thm-tactic). Previously it removed every assumption with that label. However it seemed it might sometimes be useful to have gentler behaviour; for example in porting Gappa proofs to Coq I wanted to flag certain assumptions defined by "Notation" to be automatically expanded so it was natural to use a fixed label for all of them and do REPEAT(REMOVE_THEN ...). Thu 23rd Jun 05 calc_rat.ml Inspired by porting Guillaume Marquiond's support files for Gappa, I decided it was about time I had some sort of analog to Coq's "field", so I added "REAL_FIELD". This basically uses RING after adding hypotheses "~(x = &0) ==> x * inv(x) = &1` for any inverted terms "x" appearing in the formula. But it splits up the DNF/CNF and separately tries both REAL_RING and REAL_ARITH on each part, meaning that one can sometimes deduce the nonzero-ness from some inequalities. For example it will prove this: REAL_FIELD `!c. &0 < c /\ &0 < d /\ d < b ==> (a - c) / c + (c - b) / b + ((a - c) / c) * ((c - b) / b) = (a - b) / b`;; Tue 21st Jun 05 define.ml Separated out the initial instantiation of casewise recursion (still called "instantiate_casewise_recursion") and the later first-pass processing of the "superadmissible" hypotheses, the latter now a new function "break_down_admissibility". Also reintroduced into the latter the more aggressive descent through lambdas, removing the sidecondition check on "is_applicative" for going through a lambda. Wed 15th Jun 05 sets.ml Tweaked SET_RULE so before the final MESON it also writes away "IN". This makes it better in cases where there's a mixture of set-style and predicate-style set constructs. I just threw "IN" into the same bunch of rewrites; by the top-down and minimal-generality rules this should always be fine. Fri 3rd Jun 05 Examples/binomial.ml, Examples/prime.ml Cleaned up the proofs a bit; even the definition of the binomial coefficients is easier using "define". Also added the real-number version. Also simplified the proof of EUCLID_BOUND and added an explicit "infinite set" version of the same thing: "PRIMES_INFINITE = |- INFINITE {p | prime p}". Fri 3rd Jun 05 .ml Added a recent copyright banner to all the ML files, except the "pa_j.ml" ones. Fri 3rd Jun 05 hol.ml, Examples/ Made loads keep track of files already loaded together with MD5 checksums in a variable "loaded_files". Added a "needs" function which will act like "loadt" unless the file, with unchanged checksum, has already been loaded. (It would be more rigorous to check the timestamps too, but that seems to require OS-specific stuff, as far as I can see.) Made a few uses of it to show mutual interdependencies of some of the examples. Fri 3rd Jun 05 .ocamlinit, hol.ml, sys.ml Reorganize the system a bit by removing the ".ocamlinit" file and shuffling much of the stuff from this and the root file into "sys.ml". Reorganized the root file so that simply modifying the "hol_dir" on the first line will make everything relocatable. Sun 29th May 05 Makefile Made a trivial-looking change to the Makefile: in the preprocessing directive to build the syntax extension file, changed -pp '...' to -pp "...". The double quotes also work OK in a pure Windows environment at the command prompt, whereas single quotes do not. This problem was pointed out to me by Yuri Matiyasevich. Fri 27th May 05 ind-types.ml Added "cases", which is just a slightly more convenient interface to "prove_cases_thm"; it also handles "num" specially, which otherwise fails because 0 isn't actually a constant. Fri 27th May 05 simp.ml Removed the mostly commented out stuff that's concerned with making simplifier extensions. (I kept it in "Work/simp.ml" in case I want it sometime.) Thu 18th May 05 define.ml Forced the breakdown under admissibility combinators to fail if the function being defined is used in an apparently `higher-order' way, i.e. inside a lambda or not applied to anything at one point in the body. This prevents the breakdown in difficult cases (e.g. the power series example) from going so far that the wellfoundedndess theorem is unprovable. Tue 17th May 05 meson.ml Upped the default MESON limit from 30 to 50. It's occasionally useful for goals where the proof is long but the search space is narrow, and seems harmless. When a smaller bound is required it can always be asserted. Tue 17th May 05 define.ml, wf.ml Moved the lemma MEASURE_LE back into wf.ml, since it seems more general. Mon 16th Mat 05 ind-types.ml, list.ml, sets.ml Added two new stores of injectivity and distinctness theorems and a function to simply look them up: "distinctness" and "injectivity". This avoids repeated calls and the verbose names of "prove_constructors_distinct" and "prove_constructors_injective". Took calls to those out of other files. Fri 13th May 05 iter.ml Added "left" version of the recursion equations for sums, as well as a more matchable "right" version: SUM_CLAUSES_LEFT = |- !f m n. m <= n ==> sum (m .. n) f = f m + sum (m + 1 .. n) f SUM_CLAUSES_RIGHT = |- !f m n. 0 < n /\ m <= n ==> sum (m .. n) f = sum (m .. n - 1) f + f n NSUM_CLAUSES_LEFT = |- !f m n. m <= n ==> nsum (m .. n) f = f m + nsum (m + 1 .. n) f NSUM_CLAUSES_RIGHT = |- !f m n. 0 < n /\ m <= n ==> nsum (m .. n) f = nsum (m .. n - 1) f + f n Fri 13th May 05 arith.ml, passim Also added another theorem I always end up regenerating: LT_NZ = |- !n. 0 < n <=> ~(n = 0) and changed various scripts to make use of it. Fri 13th May 05 theorems.ml, passim Added the two "distribution" theorems that otherwise I always end up recreating manually. In HOL88 they were called "LEFT_AND_OVER_OR" and "RIGHT_AND_OVER_OR"; I hope these names are more consistent with my naming conventions and a bit less of a mouthful. LEFT_OR_DISTRIB = |- !p q r. p /\ (q \/ r) <=> p /\ q \/ p /\ r RIGHT_OR_DISTRIB = |- !p q r. (p \/ q) /\ r <=> p /\ r \/ q /\ r Fixed up scripts to take out some explicit calls to TAUT replacing them with this, and likewise with IMP_IMP and IMP_CONJ. Mon 9th May 05 define.ml [new file] Added a suite of tools for making general recursive definitions. Some of the later stuff is a bit hacky, but it's a good placeholder and so long as I make things upwards compatible, we're OK. Fri 6th May 05 printer.ml Fixed an error in the printing of "let"s. Previously no brackets were printed in `(let x = 2 in x + x) = 4`. Fixed by analogy with the printing of other binders. Thu 5th May 05 int.ml Finally rewrote ARITH_RULE more carefully so that it takes out atomic terms, and generalizes them with a positivity hypothesis in the integer transition, after doing basic normalization to expand products etc. This fixes a lot of problems I've noted over the years with ARITH_RULE not assuming linearity for composite terms that are otherwise treated atomically, even including the alpha-equivalence problem: ARITH_RULE `2 * a * b <= a * b ==> a * b = 0`;; ARITH_RULE `~(k1 * k2 = 0) ==> 1 <= k1 * k2`;; ARITH_RULE `n * n * n <= n * n * n + 1 + n + n * n`;; ARITH_RULE `~(b = 0) ==> 1 <= a * n + b`;; ARITH_RULE `~(p = 0) ==> ~(binade (p,N) b = 0) ==> (binade (p,N) b + p - 1 = binade (p,N) b - 1 + p)`;; ARITH_RULE `2 EXP (p - 1) < k ==> 1 <= k`;; ARITH_RULE `2 * a * b <= a * b ==> a * b = 0`;; ARITH_RULE `2 * a * b EXP 2 <= b * a * b ==> (SUC c - SUC(a * b * b) <= c)`;; ARITH_RULE `7 <= CARD {x | x = 0} ==> 6 < CARD {x | x = 0}`;; Thu 5th May 05 canon.ml Added a conversional to apply a conversion to the "atoms" in a first-order formula: PROP_ATOM_CONV. Thu 5th May 05 wf.ml Added the more sophisticated mix of tail and wellfounded recursion: WF_REC_TAIL_GENERAL = |- !P G H. WF (<<) /\ (!f g x. (!z. z << x ==> f z = g z) ==> (P f x <=> P g x) /\ G f x = G g x /\ H f x = H g x) /\ (!f g x. (!z. z << x ==> f z = g z) ==> H f x = H g x) /\ (!f x y. P f x /\ y << G f x ==> y << x) ==> (?f. !x. f x = (if P f x then f (G f x) else H f x)) Tue 3rd May 05 wf.ml Added wellfoundedness triviality: WF_FALSE = |- WF(\x y. F). Mon 2nd May 05 grobner.ml Added conditional elimination to the pre-canonicalization in RING. So now this works: REAL_RING `p = (&3 * a1 - a2 pow 2) / &3 /\ q = (&9 * a1 * a2 - &27 * a0 - &2 * a2 pow 3) / &27 /\ x = z + a2 / &3 /\ x * w = w pow 2 - p / &3 ==> (z pow 3 + a2 * z pow 2 + a1 * z + a0 = &0 <=> if p = &0 then x pow 3 = q else (w pow 3) pow 2 - q * (w pow 3) - p pow 3 / &27 = &0)`;; Sun 24th Apr 05 sets.ml Added another "how did I manage without it for so long" result about sets and functions: CARD_LE_INJ = |- !s t. FINITE s /\ FINITE t /\ CARD s <= CARD t ==> (?f. IMAGE f s SUBSET t /\ (!x y. x IN s /\ y IN s /\ f x = f y ==> x = y)) Sat 23rd Apr 05 arith.ml, Examples/pocklington.ml Moved the definition of the "minimal" binder back into the core arithmetic file, since it seems potentially to be quite useful. Fri 22nd Apr 05 iter.ml Added HAS_SIZE_NUMSEG = |- !m n. m .. n HAS_SIZE (n + 1) - m CARD_NUMSEG_1 = |- !n. CARD(1 .. n) = n HAS_SIZE_NUMSEG_1 = |- !n. 1 .. n HAS_SIZE n Fri 22nd Apr 05 sets.ml Added CARD_PSUBSET = |- !a b. a PSUBSET b /\ FINITE b ==> CARD a < CARD b Tue 19th Apr 05 sets.ml Added IN_IMAGE into SET_TAC, which had somehow got forgotten before. Thu 14th Apr 05 sets.ml Added FORALL_IN_CLAUSES = |- (!P. (!x. x IN {} ==> P x) <=> T) /\ (!P a s. (!x. x IN a INSERT s ==> P x) <=> P a /\ (!x. x IN s ==> P x)) EXISTS_IN_CLAUSES = |- (!P. (?x. x IN {} /\ P x) <=> F) /\ (!P a s. (?x. x IN a INSERT s /\ P x) <=> P a \/ (?x. x IN s /\ P x)) Tue 12th Apr 05 iter.ml [new file] Added a new file right at the end of the core build containing segments of natural numbers together with iterated operations in general and the cases of sums of natural numbers ("nsum") and reals ("sum") in particular. Tue 12th Apr 05 realax.ml, real.ml, calc_rat.ml, Examples/analysis.ml Removed the definition of "sum" (realax.ml), as well as all the theorems about it (in real.ml) and the conversions REAL_SUM_CONV, REAL_HORNER_SUM_CONV (from calc_rat.ml). Moved it all into "Examples/analysis.ml", ready to make a nicer version the core default. Thu 7th Apr 05 meson.ml, sets.ml Added MESON, a rule form of MESON_TAC[]. Useful for generating trivial lemmas on the fly. Used it once in "sets.ml"; also made a similar tweak replacing SET_TAC[] proof with SET_RULE. Thu 7th Apr 05 bool.ml Uncommented the line giving "=" precedence 12; I'd like the keep this as the default now. Mon 4th Apr 05 printer.ml Fixed a bug where prefix operators weren't getting parenthesized in the "1000 precedence context" case; this was pointed out by Tom Hales. For example `f (~x) y` and `f x (--(&1))` were printing without brackets. Also fixed the longstanding bug of printing double negations without an intervening space. Sun 3rd Apr 05 Makefile Modified the Makefile so that "hol.multivariate" gets built from "hol" not right from scratch. Also added "hol.sosa" with SOS and analysis and "hol.card" with cardinal arithmetic. Sun 3rd Apr 05 printer.ml Fixed the printer so it doesn't put brackets in right-iterated pairs like `1,2,3`, which formerly printed as `1,(2,3)`. The problem was that when stripping an iterated pair, a comparison with the original "," was done, yet that has a different type. Fixed it by putting a special case for "," in DEST_BINARY. Sat 2nd Apr 05 sets.ml Added three theorems (formerly in Multivariate/misc.ml): SUBSET_RESTRICT = |- !s P. {x | x IN s /\ P x} SUBSET s FINITE_RESTRICT = |- !s p. FINITE s ==> FINITE {x | x IN s /\ P x} CARD_UNION_EQ = |- !s t u. FINITE u /\ (s INTER t = {}) /\ (s UNION t = u) ==> (CARD s + CARD t = CARD u) Fri 1st Apr 05 calc_rat.ml Made both REAL_RING and REAL_ARITH (hence REAL_ARITH_TAC) first write away the decimal notation "#n". Otherwise they failed on trivialities containing decimal numbers, e.g. `~(#1 = #2)`. Mon 28th Mar 05 grobner.ml, int.ml Moved back NUM_SIMPLIFY_CONV and applied it as a prenormalizer to NUM_RING; also made it rewrite SUC to something recongnizable inside. In addition, made the generic RING tactic simply ignore non-equational atoms, rather than choking on them. Fri 25th Mar 05 ind-types.ml Fixed "prove_constructors_distinct" so that it will work with num_RECURSION_STD. (The fact that `0` is not actually a constant was previously causing trouble.) Added "num" to !inductive_type_store and slightly reordered it. More dramatically, added a new store "basic_rectype_net" and a function "extend_rectype_net" to expand it, making the type definition stuff automatically add cases and distinctness theorems to it; also provided a conversion RECTYPE_EQ_CONV to apply it reasonably efficiently at depth. Since there can be quadratically many distinctness clauses, it would be better to have a clever conversion (e.g. via an auxiliary discriminator map into N) but I tried a 100-constructor type and the time to add this stuff was comparable to the time for the core definition. Thu 24th Mar 05 arith.ml Modified NUM_MULTIPLY_CONV so that it doesn't descend into subterms that are unquantified. This is analogous to what's done in CONDS_ELIM_CONV, and perhaps in fact I ought to be doing something more sophisticated like that. Anyway, this fixes a regressive bug that ARITH_RULE was failing on `2 EXP (p - 1) < k ==> 2 EXP (p - 1) <= k - 1` because it was creating a separate quantified formula in both arms of the implication. Sun 20th Mar 05 term.ml, thm.ml, basics.ml, pair.ml, drule.ml Recoded "mk_comb", "is_eq" and "dest_eq" in a somewhat more "pattern-matchy" style. Also PE'd the "mk_const" in "mk_binder". Similarly PE'd MK_FORALL, MK_EXISTS and mk_pair. The efficiency difference doesn't seem to be great, but it is a bit faster. Sun 20th Mar 05 lib.ml Made explicit definitions of "assoc" and "rev_assoc" rather than calling "find". (But kept the exception as "find" for compatibility.) Not so spectacular, but it takes a couple of percent off the load time. Sun 20th Mar 05 thm.ml Inspired by Sean's observation that "find" was taking 70% of the profile, I tried replacing "mk_eq" with a version that does an "inst" rather than a "mk_const" each time. This made the core a *lot* faster (1:25 to build versus 2:11). Mon 14th Mar 05 make.ml Added a call to "Gc.compact()" before the checkpointing in "self_destruct". This often makes the image a lot smaller. Thu 10th Mar 05 tactics.ml Removed old USE_ASSUM which I'd forgotten I ever wrote and duplicates the function of USE_THEN. Tue 8th Mar 05 parser.ml Made a little tweak to the lexical analysis so that you can use the "glue either symbolic or alphanumeric tokens together with underscores" more general in two respects: these identifiers can start with an underscore, and there can be multiple underscores in the glueing section. The immediate motivation was that I wanted to use "_|_" for a perpendicularity relation, but it seems generally sensible. Tue 8th Mar 05 grobner.ml Changed to using WEAK_DNF_CONV, which is a *lot* faster in many cases than the egregious PROP_DNF_CONV; added a custom disjunctive refuter so that we don't need to reassociate the disjuncts. I should probably go one step further as in realarith.ml and use conjuncts directly too, and hence never create the normal forms at all. Fri 4th Mar 05 grobner.ml Fixed a (fairly gross) error in the RING procedure, which was not working properly in relatively trivial cases where there were no "positive" or no "negative" elements in the certificate, notably in degenerate cases like `(x = &0) ==> (&1 = &1)`;; Wed 2nd Mar 05 sets.ml Added use of "IN_ELIM_THM" to SET_TAC. Funny this wasn't done before, but I guess I was mainly using it for combinations of the usual set operators. Also added SET_RULE, not before time. Tue 1st Mar 05 Makefile Simplified the camlp4 library search yet again using "-I +camlp4", which was pointed out to me by Trevor Morris. Thu 17th Feb 05 term.ml Slightly recoded "paconv" in what I think is a clearer style with corresponding matching of the pair of terms. Moreover it should I hope be a little more efficient, with no subcalls to "type_of" and wasteful "alphavars" checking of cases where only one term is a variable. Also did a similar simplification of "type_of"; it now doesn't have the extra "chase length" argument (which didn't pass through abstractions anyway). I hope this simpler version will also be more efficient. Mon 14th Feb 05 sets.ml Added SIMPLE_IMAGE = |- !f s. {f x | x IN s} = IMAGE f s Wed 9th Feb 05 lib.ml, pa_j_3.08.ml, pa_j_3.07+2.ml, tactics.ml, equal.ml Added a few more syntax tweaks so that new infixes work when parenthesized as in "(o)" or "(THEN)". For "o" and "upto" I removed the underscored names, which are not needed. But I can't seem to do the same for the others because the uppercaseness still spoils stuff; nevertheless the underscored names don't need to appear in sources. Also exposed the hashtable (Pa_j.ht) so that additional extensions can do the same thing without needing to copy all the pa_j stuff. I'm planning to use this to deal with Freek's Mizar Light. Wed 9th Feb 05 preterm.ml Made "make_overloadable" accept repetition of the same overload skeleton, and only fail if it differs (could even look for equivalence, i.e. mutual matchability, but I don't think it's worth the additional complexity). Wed 9th Feb 05 parser.ml Changed element-parser in set abstractions from "typed_apreterm" to a new class "nocommapreterm", which is essentially all infix terms that don't, except inside nested brackets etc., use the "," operator. This allows almost all terms to be put as elements in set abstractions, with only pairs requiring the bracketing. I think this is about the best we can do without a more radical change like using ";" or another reserved word as the separator. One might argue it would be more consistent to only allow infixes with higher precedence than ","; that would also be easy to implement if desired. Tue 8th Feb 05 grobner.ml, calc_rat.ml Backed off the introduction of IDEAL_CONV and just put in "ideal_cofactors", which returns the cofactors needed for an ideal certificate, but doesn't actually prove the equality. (This is really what I wanted for use in the Positivstellensatz prover.) However, defined appropriate REAL_IDEAL_CONV since that one might actually be useful. Mon 7th Feb 05 grobner.ml, calc_rat.ml Made a new "RING_AND_IDEAL_CONV" which returns a pair of the old RING procedure plus a more tightly restricted "IDEAL_CONV" that will prove one term is in the ideal generated by the others. In "calc_rat.ml", put instantiations of both of them to the reals: REAL_RING and REAL_IDEAL_CONV. Sun 6th Feb 05 printer.ml Added a flag "print_all_thm", which suppresses the printing of theorem hypotheses if set to false. Sean McLaughlin was wanting this, and I sometimes find it useful myself. It was in fact the default in old HOL88 not to print hypotheses. Sat 5th Feb 05 Makefile, make.ml At the suggestion of Freek Wiedijk, changed the checkpointing signal from TSTP to USR1. This means you can still do the usual ctrl-z suspend of a checkpointed process without causing another checkpoint dump. Thu 3rd Feb 05 arith.ml Added natural number counterparts of the real "without loss of generality" lemmas: WLOG_LE = |- (!m n. P m n <=> P n m) /\ (!m n. m <= n ==> P m n) ==> (!m n. P m n) WLOG_LT = |- (!m. P m m) /\ (!m n. P m n <=> P n m) /\ (!m n. m < n ==> P m n) ==> (!m y. P m y) Thu 3rd Feb 05 canon.ml Refined CONDS_ELIM_CONV/CONDS_ELIM_CONV' so that they appropriately share conditionals with a common condition, even if the things decided by those conditions are different. For example, (if x = 1 then a = 1 else b = 2) /\ (if x = 1 then T else F) will now generate just one case split over x = 1, not two nested ones. Although this may seem like a relatively minor efficiency issue, the change to a more "global" CONDS_ELIM_CONV yesterday was actually breaking some proofs, because the redundant case splits were then causing MESON to hit its split limit. Thu 3rd Feb 05 int.ml Slightly tweaked NUM_SIMPLIFY_CONV to apply NUM_REDUCE_CONV before the elimination of division etc. This makes ARITH_RULE a bit better at handling cases involving unevaluated constant expressions. Wed 2nd Feb 05 canon.ml, realarith.ml, int.ml Modified CONDS_ELIM_CONV so that it takes the topmost formula not involving quantifiers and does the main transformation there; on balance this seems to give the best chance of sharing several with identical tests, and will also make the parity of the newly introduced formulas more controllable; the latter was a problem before in REAL_ARITH. Added parallel CONDS_ELIM_CONV' for contexts where you want a disjunction; attempted to make both keep the right parity as they descend terms (but don't expand any <=>s so you can't really do it right there). Used that in NUM_SIMPLIFY_CONV, which makes conditional-elimination in the discrete decision procedures better. Even though I can now rely better on parity, I added the tweak in REAL_ARITH of doing a preliminary NNF conversion and applying CONDS_ELIM_CONV over the disjuncts, to avoid gratuitous "sharing" among separate subgoals. Fixed a silly bug in INT_ARITH, which simply wasn't applying the initialization conversion and so wouldn't always handle conditionals etc. properly; took the chance to make it more like the reals one. Seems to work pretty well, and this fixes a few regressions such as this (previously this generated a complicated DNF with a = b and ~(a = b) in one disjunct, which wasn't picked up by the core routine). REAL_ARITH `~(x - (if a = b then &1 else &0) <= &0) ==> ~(x = &0)`;; ARITH_RULE `~(multiplicity M x - (if x = a then 1 else 0) = 0) ==> ~(multiplicity M x = 0)` Wed 2nd Feb 05 tactics.ml Only 10 years too late, introduced a few tactics for making more convenient use of labelled assumptions: USE_THEN, REMOVE_THEN and SUBGOAL_TAC. I was actually doing a proof today where quoting the terms was becoming too painful. Also added REPLICATE_TAC, which I occasionally want and end up hacking round. Tue 1st Feb 05 grobner.ml Added "grobner_ideal" to certify that a polynomial is in the ideal generated by another; via a tweak to "resolve_proof" so Start(-1) is treated as zero, I could essentially use the same framework. Next I'll integrate it into an actual HOL conversion. Tue 1st Feb 05 realarith.ml Made the initial normalization, once it's reached a formula (negating the thing to be proved for refutation) ?x1..xn. !y1..ym.... no longer just die if m =/= 0. This spoiled the "ignoring" of hopefully irrelevant subterms, as in: REAL_ARITH `(!x. x + 1 = SUC x) ==> (y:real = y)`;; Now it just specializes them, so it not only does that but even will use universal formulas provided that a completely arbitrary instantiation works: REAL_ARITH `(!x:real. x < x) ==> (w = z:real)`;; This ought to propagate to integer and natural versions too and so fix one or two regressions. Mon 31st Jan 05 grobner.ml Fixed a little potential performance bug: the initial critical pairs were not sorted by "forder", so that entire optimization was only partly useful. Fri 28th Jan 05 Makefile Tidied up the Makefile a bit with more verbose explanations, and made use of "camlp4 -where" instead of the former hack to find the camlp4 library directory. Also added a dependency on "pa_j_3.08.2" (these flaws were pointed out by Tom Hales). Mon 24th Jan 05 bool.ml Backed off the change to precedence of "=" a while longer since I want to make a preliminary quasi-release with Tom's Flyspeck proof. Fri 21st Jan 05 tactics.ml Made the change to ABBREV_TAC slightly less sharp by not looking in the "assumptions of the assumptions" for existing variables. Thu 20th Jan 05 int.ml Made a few little tweaks and improvements to INT_ARITH: treated abs (and now max and min) using the same bubble treatment rather than eliminating them, and avoided explicitly locating and replacing the alien subterms, leaving that to REAL_ARITH; this makes the basic approach applicable even in situations with a more complicated quantifier structure (in case we ever care). Thu 20th Jan 05 bool.ml, passim Finished updating all the Examples for the new precedences; hence made it the default. Now I just need to fix up some of my own proofs, which I can do little by little. Thu 20th Jan 05 realarith.ml, real.ml Moved REAL_POS back into "realarith.ml" for use in REAL_ARITH. Then modified REAL_ARITH (not the general wrapper REAL_ARITH, since one really needs a different notion of 'variable' in the nonlinear case) so that it identifies "alien" subterms of the form "&t" for t not a numeral, and adds positivity hypotheses. This means that REAL_ARITH `&k + x:real <= x`, and more interesting examples, work. But the real advantage is that alien subterms in ARITH_RULE should inherit the effects. Thu 20th Jan 05 int.ml Added a conversion NUM_TO_INT_CONV to convert an assertion over N to a corresponding one over Z. Previously this was embedded in INT_ARITH itself and would only handle universal formulas. But now I want to use it as a subcomponent in Cooper's algorithm too. Modified ARITH_RULE to use this rather than the previous stuff. This now no longer does anything about alien subterms, i.e. "&t" for t a non-numeral. I plan to put this into REAL_ARITH (and so implicitly INT_ARITH) instead, since it may occasionally be useful there anyway. On a quick test of explicitly grepped cases, the only failure in the new version that worked before was the following, and since the3 solution (rewrite with sub-distributivity) seems ad hoc I changed the proof to remove it. `n * b <= 1 * b + c ==> (n - 1) * b <= c` The new version also fixes a number of the bugs I had against the old ARITH_RULE, e.g. `p - 1 <= 2 * p` and `g (f m) < f m + g (f m) + 1`. It still doesn't fix some cases where the bubbling through `known' operators will fail to leave a simple "alien" term for REAL_ARITH, e.g. `~(k1 * k2 = 0) ==> 1 <= k1 * k2`. But they're a bit marginal and maybe I should defer them to a more interesting nonlinear case anyway. Wed 19th Jan 05 int.ml Added NUM_SIMPLIFY_CONV, a somewhat more capable conversion for initial canonicalization of natural number formulas with arbitrary quantifier structure, which reduces pretty much everything to addition and multiplication. It does take care to keep the result universal if the original was, and returns something in NNF. Eventually this will be used in ARITH_RULE (and in Cooper's procedure in the examples). Wed 19th Jan 05 arith.ml Slightly modified NUM_MULTIPLY_CONV to make it a bit better as a component elsewhere: it now puts the newly introduced subterm in NNF, so if the original was in NNF so will be the result. Added a flag which if set will stop it from descending through negations (hence better for a term already in NNF to keep the signs of the subsidiary quantifiers straight). I also removed the initial LAMBDA_ELIM_CONV and COND_ELIM_CONVs which I'd rather be able to choreograph separately. Sun 16th Jan 05 passim Finished fixing up the core so that it works with new scheme for "<=>", and for now returned the precedence of "=" to 2. Sat 15th Jan 05 bool.ml and passim Started the project of making the core safe with "<=>" and a higher precision for "=". Started by adding the following to "bool.ml": parse_as_infix("<=>",(2,"right"));; override_interface ("<=>",`(=):bool->bool->bool`);; parse_as_infix("=",(12,"right"));; Then went through changing "=" into "<=>" in most places; at least all places where it would otherwise break, though I probably missed quite a few "harmless" ones. The medium-term goal is to make the entire core, libraries and proofs I care about work equally well with or without the last one of the above three lines. Then I'll probably be brave enough to use the last line in the production system, and start relying on the new precedence. At least I'm going to use the first two. Fri 14th Jan 05 calc_rat.ml Fixed REAL_RAT_INV_CONV so that "inv(&1 / &n)" or "inv(-- &1 / &n)" return integer constants not non-canonical things like "&n / &1". Fri 14th Jan 05 type.ml, term.ml Added "types()" and "constants()" to return the type and term constants. As Freek had pointed out, now the underlying lists are hidden, one wants some way of looking at them. Fri 14th Jan 05 basics.ml Deleted the near-duplicate definition of "vfree_in", which is already defined in "term.ml". They are perfectly identical assuming the first argument is indeed a variable; there is a small difference in that the term.ml version will accept a constant and tell you if it's free, whereas the deleted one will give false if the first argument is not a variable. I don't think this is ever relied on. I could change it, but (a) it could actually be useful, and (b) it might be marginally faster, and (c) it keeps the core one line shorter. Fri 14th Jan 05 tactics.ml Modified ABBREV_TAC so that it will fail if the chosen abbreviating variable is already used somewhere in the goal. Previously it would just silently pick a variant. Fri 14th Jan 05 printer.ml Made the printer print a space between two "--" negation symbols; previously it wasn't doing so and the result "----" didn't re-parse. Also slightly generalized the circumstances under which it adds a space in "-- &n": now it does it even if n is not a numeral. Fri 14th Jan 05 int.ml Added new definitions int_max and int_min together with int_max_th and int_min_th, and updated the theorem-lifter to deal with them. Hence added integer clones of all the real theorems I added in the last batch: INT_MIN_MAX = |- !x y. min x y = --(max (--x) (--y)) INT_MAX_MIN = |- !x y. max x y = --(min (--x) (--y)) INT_MAX_MAX = |- !x y. x <= max x y /\ y <= max x y INT_MIN_MIN = |- !x y. min x y <= x /\ min x y <= y INT_MAX_SYM = |- !x y. max x y = max y x INT_MIN_SYM = |- !x y. min x y = min y x INT_LE_MAX = |- !x y z. z <= max x y = z <= x \/ z <= y INT_LE_MIN = |- !x y z. z <= min x y = z <= x /\ z <= y INT_LT_MAX = |- !x y z. z < max x y = z < x \/ z < y INT_LT_MIN = |- !x y z. z < min x y = z < x /\ z < y INT_MAX_LE = |- !x y z. max x y <= z = x <= z /\ y <= z INT_MIN_LE = |- !x y z. min x y <= z = x <= z \/ y <= z INT_MAX_LT = |- !x y z. max x y < z = x < z /\ y < z INT_MIN_LT = |- !x y z. min x y < z = x < z \/ y < z INT_MAX_ASSOC = |- !x y z. max x (max y z) = max (max x y) z INT_MIN_ASSOC = |- !x y z. min x (min y z) = min (min x y) z INT_MAX_ACI = |- (max x y = max y x) /\ (max (max x y) z = max x (max y z)) /\ (max x (max y z) = max y (max x z)) /\ (max x x = x) /\ (max x (max x y) = max x y) INT_MIN_ACI = |- (min x y = min y x) /\ (min (min x y) z = min x (min y z)) /\ (min x (min y z) = min y (min x z)) /\ (min x x = x) /\ (min x (min x y) = min x y) Fri 14th Jan 05 list.ml Replaced EX_MEM with the equation the other way round. It's never used anywhere anyway and it seems better to have it consistent with ALL_MEM. It's now: EX_MEM = |- !P l. (?x. P x /\ MEM x l) = EX P l Thu 13th Jan 05 Examples/pratt.ml, Examples/pocklington.ml Renamed an overwritten theorem CONG_SUB_CASES. Actually it's only in pocklington that it gets overwritten, but it seems wise to be consistent. Thu 13th Jan 05 simp.ml, class.ml, pair.ml, wf.ml, ind-types.ml, realax.ml, sets.ml Finally cured the minor but persistent irritant that one can't easily use ETA_AX as a rewrite because it will in general loop on any lambda owing to the higher-order match. The fix was to modify "net_of_thm" to treat an (unconditional) rewrite alpha-equivalent to ETA_AX as a special case with a first order match. Changed various proof scripts to avoid ETA_CONV and just use ETA_AX as a rewrite, which is the usage pattern I now want to establish. Thu 13th Jan 05 real.ml Added a large suite of basic lemmas about max and min, all of which are proved automatically by the new REAL_ARITH_TAC. REAL_MIN_MAX = |- !x y. min x y = --(max (--x) (--y)) REAL_MAX_MIN = |- !x y. max x y = --(min (--x) (--y)) REAL_MAX_MAX = |- !x y. x <= max x y /\ y <= max x y REAL_MIN_MIN = |- !x y. min x y <= x /\ min x y <= y REAL_MAX_SYM = |- !x y. max x y = max y x REAL_MIN_SYM = |- !x y. min x y = min y x REAL_LE_MAX = |- !x y z. z <= max x y = z <= x \/ z <= y REAL_LE_MIN = |- !x y z. z <= min x y = z <= x /\ z <= y REAL_LT_MAX = |- !x y z. z < max x y = z < x \/ z < y REAL_LT_MIN = |- !x y z. z < min x y = z < x /\ z < y REAL_MAX_LE = |- !x y z. max x y <= z = x <= z /\ y <= z REAL_MIN_LE = |- !x y z. min x y <= z = x <= z \/ y <= z REAL_MAX_LT = |- !x y z. max x y < z = x < z /\ y < z REAL_MIN_LT = |- !x y z. min x y < z = x < z \/ y < z REAL_MAX_ASSOC = |- !x y z. max x (max y z) = max (max x y) z REAL_MIN_ASSOC = |- !x y z. min x (min y z) = min (min x y) z REAL_MAX_ACI = |- (max x y = max y x) /\ (max (max x y) z = max x (max y z)) /\ (max x (max y z) = max y (max x z)) /\ (max x x = x) /\ (max x (max x y) = max x y) REAL_MIN_ACI = |- (min x y = min y x) /\ (min (min x y) z = min x (min y z)) /\ (min x (min y z) = min y (min x z)) /\ (min x x = x) /\ (min x (min x y) = min x y) Mon 10th Jan 05 realarith.ml Fixed a little bug: the translator for "Square" in Positivstellensatz proofs was not forcing normalization. Of course this only matters for the extra stuff I'm doing on nonlinear arithmetic outside the core. Mon 10th Jan 05 hol.ml Added an explicit printer for the type "num". I thought this was already there, but apparently not; I guess I just mostly look at small numbers. Mon 10th Jan 05 lib.ml, arith.ml, basics.ml, calc_int.ml, grobner.ml, normalizer.ml, preterm.ml, realarith.ml Added a few bits and pieces: num_0, num_1, num_2, num_10, pow2, pow10, increasing, decreasing. Swept through the other files changing "Int x" to "num_x" for x in 0, 1, 2 and 10. This is hardly a big issue, but we might as well get a little extra partial evaluation. Sat 8th Jan 05 realarith.ml Fixed yet another tiny failure in REAL_ARITH: it was failing if the initial canonicalization already returned `F`. Fri 7th Jan 05 calc_int.ml Made the destructor and tester functions for rationals an integers more careful about excluding non-canonical cases like '--(&0)`, `&5 / &1` and `&2 / &4`. This is more important now these are used in REAL_ARITH as the justification for doing no internal simplification. Fri 7th Jan 05 basics.ml Modified some of the destructors to use direct pattern-matching rather than nested primitive destructors. Rather disappointingly, I see no speedup; in fact if anything it's slower. Thu 6th Jan 05 canon.ml Added two more carefully implemented variants of CNF and DNF, though kept the old olds since they work inside a quantifier prefix and that's depended on sometimes. WEAK_CNF_CONV and WEAK_DNF_CONV force the appropriate conj-of-disj or vice versa, but no association, and STRONG_CNF_CONV and STRONG_DNF_CONV go further and AC-canonicalize. Perhaps I should modify the full conversions to be comparably efficient; at the moment they're just crude rewrites and may be very slow. Thu 6th Jan 05 calc_int.ml, realarith.ml, real.ml, calc_rat.ml Radically updated the reals decision procedure REAL_ARITH (and so REAL_ARITH_TAC). The new version copes with arbitrary numbers, rationals and powers, and uses better data structures for normalization so it can be dramatically faster on big algebraic simplification tasks (though the more general approach can make it a little slower on smaller problems). In addition, sepearated out the integer calculations into a (new) separate file "calc_int.ml", and the decision procedure into another new file "realarith.ml". The procedure is highly parametrized, and in fact it's first bootstrapped over the integers and only extended to cope with rationals later in "calc_rat.ml". It's also ready to incorporate the kind of more general Positivstellensatz certificates that nonlinear procedures may want to use. Thu 6th Jan 05 grobner.ml, list.ml Renamed the more general basis-computing function "grobner_basis"; it seemed a bit inelegant to use "grobner" twice for different things. Removed duplicate (except for one quantifier; anyway it was unused) FORALL_ALL. Thu 6th Jan 05 printer.ml, basics.ml, bool.ml Removed some duplicate definitions of syntax operations; moved all those possible (basically, testers and destructors but not constructors) from "bool.ml" back to "basics.ml", even when the underlying constants haven't been defined. Thu 6th Jan 05 normalizer.ml Fixed a little bug in the normalizer, which would choke on "x pow y" for "y" a non-numeral, rather than the desirable behaviour of returning a reflexive theorem in such a case. Wed 5th Jan 05 realax.ml, int.ml Made the natural-number injection symbol "&" overloaded rather than separately overriden for the integers and reals. Wed 5th Jan 05 calc_rat.ml Fixed "dest_ratconst" and "is_ratconst" (and hence "rat_of_term" which uses the former) so that they refuse to accept rationals with negative or zero denominators. Tue 4th Jan 05 canon.ml, meson.ml Installed a new version of NNF_CONV. This provides a new more general GEN_NNF_CONV which allows a user-installed conversion for the atomic formulas. Moreover, it performs a fancier recursion (computing the NNF of a formula and its negation in parallel) to avoid some recomputation when doing multiple splits of "iff". Finally, it's just more carefully (and painfully) coded than the previous version, which was simply done by rewrites. Threw away the CNF-angling "NNFC_CONV" and made this standard function enter a similar mode below universal quantifiers; hence used the new function in MESON_TAC. This should be better than the old version of NNFC_CONV anyway, which stupidly did the same thing even before passing through quantifiers, less good for initial case splits. Indeed, the new version only gives 1024 basic MESON problems when building the core, whereas the old gives 1939. The speedup is actually quite significant: according to some quick tests, the time to build the core on my laptop is reduced by about 7% to just under 2 minutes. Tue 4th Jan 05 theorems.ml Added EXISTS_UNIQUE = |- !P. (?!x. P x) = (?x. P x /\ (!y. P y ==> (y = x))) Tue 28th Dec 04 lib.ml Incorporated my entire "fpf.ml" file for finite partial functions, which notably adds the "combine" function to what I had before. Tue 28th Dec 04 realax.ml Added definitions of "real_max" and "real_min", as well as incorporating their overloaded forms "max" and "min". Mon 27th Dec 04 meson.ml Deleted unused PREMESON_CANON_TAC. Fri 17th Dec 04 meson.ml Started reworking the basics of MESON_TAC's first order core. Introduced somewhat different and simpler unification algorithm that keeps instantiation in a "graph" (not fully solved) form; also changed the equality test function. Made the unification function more rationally keep the two terms asymmetrical so that you know which the variable offset applies to. Well, this is now a bit neater, but outside the individual unification steps I still in general solve the graph; eventually I'd like to fix this too, and move to a more efficient data structure for instantiations, but that was enough for one day. Mon 13th Dec 04 sets.ml Added some theorems about the cardinality of function spaces with finite "support"/"domain": HAS_SIZE_FUNSPACE = |- !d n t m s. s HAS_SIZE m /\ t HAS_SIZE n ==> {f | (!x. x IN s ==> f x IN t) /\ (!x. ~(x IN s) ==> (f x = d))} HAS_SIZE (n EXP m) CARD_FUNSPACE = |- !s t. FINITE s /\ FINITE t ==> (CARD {f | (!x. x IN s ==> f x IN t) /\ (!x. ~(x IN s) ==> (f x = d))} = CARD t EXP CARD s) FINITE_FUNSPACE = |- !s t. FINITE s /\ FINITE t ==> FINITE {f | (!x. x IN s ==> f x IN t) /\ (!x. ~(x IN s) ==> (f x = d))} Changed the proofs of HAS_SIZE_POWERSET and FINITE_POWERSET to use that as a lemma, and added CARD_POWERSET = |- !s. FINITE s ==> (CARD {t | t SUBSET s} = 2 EXP CARD s) Mon 13th Dec 04 sets.ml, preterm.ml, printer.ml Fixed a long-standing irritation, that IN_ELIM_THM would disturb the internal structure of set abstractions that it didn't eliminate; and at the same time some other apparently harmless theorems like CONJ_ACI could do the same. Fixed this by defining a new constant SETSPEC and using this instead of the conjunction and equation it's equivalent to. SETSPEC = |- SETSPEC v P t = P /\ (v = t) Correspondingly changed IN_ELIM_THM to be more delicate; note that we use a high-order rewrite to ensure that we only eliminate SETSPEC as part of a successful overall elimination: IN_ELIM_THM = |- (!P x. x IN GSPEC (\v. P (SETSPEC v)) = P (\p t. p /\ (x = t))) /\ (!p x. x IN {y | p y} = p x) /\ (!P x. GSPEC (\v. P (SETSPEC v)) x = P (\p t. p /\ (x = t))) /\ (!p x. {y | p y} x = p x) /\ (!p x. x IN (\y. p y) = p x) More ambitiously, I reverted to an old policy of making the implicit bound variables in a set abstraction the ones free on both sides of the "|" in a set enumeration. However, to avoid the problems that caused me to abandon this policy originally, I special-case the two cases: (1) when there is exactly one free variable on the left of "|", we always consider it bound, and (2) if there are no free variables on the right of the "|" we just consider all those on the left bound. I hope this will be much more useful. Fri 10th Dec 04 sets.ml Added a generalized form of yesterday's theorem; the original is a trivial consequence of this one. SURJECTIVE_IFF_INJECTIVE_GEN = |- !s t f. FINITE s /\ FINITE t /\ (CARD s = CARD t) /\ IMAGE f s SUBSET t ==> ((!y. y IN t ==> (?x. x IN s /\ (f x = y))) = (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y))) Also added a couple of simple consequences: IMAGE_IMP_INJECTIVE_GEN = |- !s t f. FINITE s /\ (CARD s = CARD t) /\ (IMAGE f s = t) ==> (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y)) IMAGE_IMP_INJECTIVE = |- !s f. FINITE s /\ (IMAGE f s = s) ==> (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y)) Also added one more triviality: HAS_SIZE_CARD = |- !s n. s HAS_SIZE n ==> (CARD s = n) Thu 9th Dec 04 sets.ml Added the classic theorem that a function from a finite set into itself is injective iff surjective. Quite surprising I've managed without it all these years... SURJECTIVE_IFF_INJECTIVE = |- !s f. FINITE s /\ IMAGE f s SUBSET s ==> ((!y. y IN s ==> (?x. x IN s /\ (f x = y))) = (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y))) Mon 6th Dec 04 README, Examples/pratt.ml, Examples/pocklington.ml Some minor tidying up: described the new Make process better and made the two primality-proving routines use the standard "Filename.temp_file" function rather than an ad-hoc alternative of my own. Fri 3rd Dec 04 sets.ml Added HAS_SIZE_CLAUSES = |- (s HAS_SIZE 0 = (s = {})) /\ (s HAS_SIZE SUC n = (?a t. t HAS_SIZE n /\ ~(a IN t) /\ (s = a INSERT t))) which is more useful in many cases than the "more general" HAS_SIZE_SUC; in particular one can just blindly apply num_CONV and that rewrite rule to produce expressions for a set known to have a particular number of elements. Added a conversion HAS_SIZE_CONV to do this expansion in a fairly efficient and controlled way. Also added PAIRWISE and "pairwise", which I hope will be useful. Thu 2nd Dec 04 Makefile Tidied up the Makefile a bit and made it issue a more explanatory failure if the native OS isn't Linux and the user tries to build a checkpointed image. Wed 1st Dec 04 Makefile, make.ml [new file] Set up a build process for images using "ckpt". It's slightly hacky with the process sending out a signal to suspend itself, but it seems to work. Wed 1st Dec 04 lib.ml Changed the representation of finite partial functions to the canonical Patricia tree form instead of the AVL variant that was there before. Wed 1st Dec 04 canon.ml, ind-defs.ml, ind-types.ml, Examples/transc.ml Added hand-coded rules CONJ_ACI_RULE and DISJ_ACI_RULE for reordering conjunctions and disjunctions. Got rid of some more ad-hoc implementations in other files. Tue 30th Nov 04 sets.ml Added CARD_EQ_0 = |- !s. FINITE s ==> ((CARD s = 0) = (s = {})) and IMAGE_CONST = |- !s c. IMAGE (\x. c) s = if s = {} then {} else {c} Mon 29th Nov 04 sets.ml Modified FORALL_PASTECART to |- (!p. P p) = (!x y. P (pastecart x y)) and added EXISTS_PASTECART = |- (?p. P p) = (?x y. P (pastecart x y)) Sat 27th Nov 04 printer.ml Fixed the internal precedence setting when printing lists, so that no redundant bracketing is done. Before `[1,2]` would print as `[(1,2)]`. Of course some people might regard that as a feature... Sat 27th Nov 04 sets.ml Added a natural dual to FORALL_IN_IMAGE, namely EXISTS_IN_IMAGE = |- !f s. (?y. y IN IMAGE f s /\ P y) = (?x. x IN s /\ P (f x)) as well as this useful lemma: SUBSET_IMAGE = |- !f s t. s SUBSET IMAGE f t = (?u. u SUBSET t /\ (s = IMAGE f u)) Replaced the old FINITE_SUBSET_IMAGE, which is now called FINITE_SUBSET_IMAGE_IMP, with the stronger result: FINITE_SUBSET_IMAGE = |- !f s t. FINITE t /\ t SUBSET IMAGE f s = (?s'. FINITE s' /\ s' SUBSET s /\ (t = IMAGE f s')) Wed 24th Nov 04 list.ml Added LENGTH_MAP2 = |- !f l m. (LENGTH l = LENGTH m) ==> (LENGTH (MAP2 f l m) = LENGTH m) Tue 23rd Nov 04 list.ml Added LENGTH_EQ_CONS = |- !l n. (LENGTH l = SUC n) = (?h t. (l = CONS h t) /\ (LENGTH t = n)) Fri 19th Nov 04 sets.ml Added two more useful theorems of function calculus, generalizing the "injectivity" and "surjectivity" theorems, which are the special case of identity: FUNCTION_FACTORS_LEFT = |- !f g. (!x y. (g x = g y) ==> (f x = f y)) = (?h. f = h o g) FUNCTION_FACTORS_RIGHT = |- !f g. (!x. ?y. g y = f x) = (?h. f = g o h) I suppose I ought, by analogy with the injectivity and surjectivity ones, to prove them with set restrictions too. Mon 15th Nov 04 real.ml Added REAL_SUB_INV = |- !x y. ~(x = &0) /\ ~(y = &0) ==> (inv(x) - inv(y) = (y - x) / (x * y)) Mon 15th Nov 04 sets.ml Added the suite of useful lemmas relating (local) injectivity and surjectivity to left and right inverses: SURJECTIVE_ON_RIGHT_INVERSE = |- !f t. (!y. y IN t ==> (?x. x IN s /\ (f x = y))) = (?g. !y. y IN t ==> g y IN s /\ (f (g y) = y)) INJECTIVE_ON_LEFT_INVERSE = |- !f s. (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y)) = (?g. !x. x IN s ==> (g (f x) = x)) SURJECTIVE_RIGHT_INVERSE = |- (!y. ?x. f x = y) = (?g. !y. f (g y) = y) INJECTIVE_LEFT_INVERSE = |- (!x y. (f x = f y) ==> (x = y)) = (?g. !x. g (f x) = x) Sun 14th Nov 04 sets.ml Added FORALL_PASTECART = |- (!p. P (fstcart p) (sndcart p)) = (!x y. P x y) Mon 8th Nov 04 sets.ml Added FORALL_IN_UNIONS = |- !P s. (!x. x IN UNIONS s ==> P x) = (!t x. t IN s /\ x IN t ==> P x) Fri 5th Nov 04 sets.ml Added PASTECART_EQ = |- !x y. (x = y) = (fstcart x = fstcart y) /\ (sndcart x = sndcart y) Fri 5th Nov 04 sets.ml Added DIMINDEX_GE_1 = |- !s:A->bool. 1 <= dimindex(s), which is trivial but often useful. Thu 4th Nov 04 sets.ml Added stuff for pasting together two Cartesian products; the type constructor "finite_sum", pairing and projection functions "pastecart", "fstcart" and "sndcart", together with the following theorems: DIMINDEX_HAS_SIZE_FINITE_SUM, DIMINDEX_FINITE_SUM, FSTCART_PASTECART, SNDCART_PASTECART and PASTECART_FST_SND. Thu 4th Nov 04 real.ml, int.ml, Examples/analysis.ml Fixed a trivial and long-standing bug. The theorem REAL_LET_ANTISYM is what should have been called REAL_LTE_ANTISYM. Moreover, the latter didn't exist but was misspelled as REAL_LTE_ANTSYM, and was the wrong way round too! This has been changed to do the obviously right thing. The corresponding change was made for the integers too. Thu 4th Nov 04 parser.ml Added the "atleast" parser combinator, just because I've found it handy in other contexts. Tue 2nd Nov 04 sets.ml Added UNION_SUBSET = |- !s t u. (s UNION t) SUBSET u = s SUBSET u /\ t SUBSET u Thu 28th Oct 04 type.ml, term.ml After asking on the OCaml list, I found out (thanks to Andrej Bauer, William Lovas and Brian Rogoff) that there's a way of making type constructors visible in pattern matching yet still disallowing their use as constructors, by using the explicit type definition together with the "private" keyword in the signature. So I changed the abstract type definitions of types and terms in that way. That's good, because it won't require people like Tom Hales who've written code by pattern-matching to change it now I have a proper LCF core. And maybe I can start defining things that way myself now; it should be much cleaner and probably more efficient. Tue 19th Oct 04 Makefile Modified the Makefile to eliminate the hardwired path to camlp4 (this was pointed out by Sean); now it looks for the ocaml binary and guesses the camlp4's library directory based on that. Tue 12th Oct 04 type.ml, term.ml, thm.ml Changed the three instances of "open" for the core modules to "include". This was pointed out by Freek --- the effect is the same but one no longer gets the irritating qualifiers "Term.term" etc. Mon 11th Oct 04 arith.ml Added NUM_MULTIPLY_CONV, which should fairly reliably remove all the "nasties" from statements about N: predecessor, cutoff subtraction, division and modulus, plus abstractions and conditionals. This could possibly be a precursor to a clever decision procedure; at least it makes it easy to adapt my Cooper implementation for the linear case. Mon 11th Oct 04 arith.ml Modified the definitions of DIV and MOD to force them both to be zero when the divisor is zero. This does make DIV consistent with the reals in that case, and more importantly it made possible the next part of the change: changed DIVMOD_ELIM_THM into a simple equality. Previous fixes were on the plane; made the last in O'Hare airport. Mon 11th Oct 04 meson.ml Moved the generalization over free variables in MESON to a later phase when everything has been loaded into the conclusion. This works better in situations where you have a given Boolean variable free in both assumptions and conclusion, by eventually forcing a case split. Mon 11th Oct 04 printer.ml Modified the printing of types so it uses precedence and associativity of type constructors in order to reduce the bracketing; formerly it just blindly bracketed all nested infixes. Mon 11th Oct 04 printer.ml Fixed a misfeature in the treatment of binary operators. They are broken up iteratively as an easy was of getting a list to then treat according to specified associativity. But previously the splitup was done based only on the head operator name, ignoring the type, so that the apparent iteration of a single operator could actually be multiple type instances. This was particularly striking after doing the <=>/= redefinition below, when `1 = 2 <=> 2 = 3` would print as `1 = 2 <=> 2 <=> 3`;; Sat 9th Oct 04 grobner.ml [new file] Added a new "generic" Grobner basis procedure and its instantiation to N. This essentially uses Nullstellensatz certificates, so it's only really complete for C, but still gets a lot of "natural" stuff on other rings and semirings. It uses Strong Nullstellensatz certificates to avoid using any field properties to do the Rabinowitsch trick. (If one does have a field it's probably more efficient to use Weak Nullstellensatz certificates directly and in fact do proof generation without explicitly contructing the certificate --- the hooks are all there to do this if desired.) Roughly, it's complete for the universal theory of commutative cancellation semirings with no nilpotents and characteristic zero, I think. Sat 9th Oct 04 lib.ml, calc_rat.ml Added "lcm_num", "merge" and "mergesort", and moved "gcd_num", "numdom", "numerator" and "denominator" back from "calc_rat.ml". Also modified "allpairs" to use appending rather than unioning; I think in most applications that's actually what's wanted. Sat 9th Oct 04 parser.ml Removed the hacky special treatment of equality among the binary operators. I think this was only necessary because of the lack of a left bracketing symbol in the old conditionals notation, which I've just scrubbed. Now the infixes are treated in a completely regular way based only on the list of precedences and associativities. Among other things, this makes possible a simple change to a "better" way of dealing with "=" and "<=>" that Freek at least wanted and I might end up making the default: parse_as_infix("<=>",(2,"right"));; override_interface ("<=>",`(=):bool->bool->bool`);; parse_as_infix("=",(15,"right"));; Sat 9th Oct 04 parser.ml, Examples/analysis.ml, Examples/transc.ml, Examples/poly.ml Eliminated the old notation ".. => .. | .." for conditionals; now only the preferred notation "if .. then .. else .." is accepted. Removed it from the parser and eliminated the only remaining uses elsewhere in the system. Also unreserved the "=>" symbol, though not "|" which is also special in set abstractions. Sat 9th Oct 04 term.ml Changed "aconv" to use "Pervasives.compare tm1 tm2 = 0" rather than simply "tm1 = tm2". This is Xavier Leroy's recommended solution in response to an observation by Christophe Raffalli on the OCaml list that the builtin equality no longer guarantees an early-out in the case of pointer eq. A quick test on my laptop indicates that the build time for the core goes down from 2m7.845s to 1m58.681, nearly an 8% improvement. With non-trivial proofs containing big terms, the difference might be much more. Wed 6th Oct 04 drule.ml Modified "PART_MATCH" so that it fails if any type variables in the hypothesis get instantiated; previously it only checked for terms in the assumptions becoming instantiated, and this can lead to some oddities in rewriting, e.g. let th1 = ASSUME `f = \s f x:A. if x IN s then f(x) else &0`;; let th2 = BETA_RULE (AP_THM th1 `s:A->bool`);; let th3 = SYM(BETA_RULE(AP_THM th2 `f:A->real`));; GEN_REWRITE_CONV I [th3] `\a:B. if a IN s then f a else &0`;; Tue 5th Oct 04 meson.ml Made one more small tweak to the new MESON canonicalization, to generalize the goal first. Though this normally disappears, it's beneficial if the variables are Boolean since it allows for subsequent elimination. This finally fixes a problem Mike Gordon pointed out back on 15th March 2001, that MESON_TAC fails on `?b. b = ~a`. Mon 4th Oct 04 arith.ml, real.ml, int.ml, Examples/pratt.ml Examples/pocklington.ml Examples/cong.ml Introduced congruence notation "(x == y)" and three overloaded variants of "mod" for the three basic number systems. All non-trivial stuff is delegated to Examples files, but at least this gives a uniform approach; formerly the two prime-number-proving files used a congruence that, while superficially similar, was defined in a completely different way. Mon 4th Oct 04 meson.ml With some trepidation, put a new initial normalizer into MESON_TAC. This should split conditionals and treat select-terms and abstractions with a degree of intelligence; this has long been a weak spot of MESON_TAC. I was pleasantly surprised how little stuff broke, and the ability to handle more things is already quite useful. Mon 4th Oct 04 bool.ml Added syntax operations for "unique exists": is_uexists, dest_uexists and mk_uexists. Mon 4th Oct 04 class.ml Added FORALL_BOOL_THM and EXISTS_BOOL_THM to deal with quantification over Booleans. Mon 4th Oct 04 canon.ml Eliminated the old "EQ_ABS_CONV" and "DELAMB_CONV", which are respectively not so very useful and only used by MESON. Slightly modified SPLIT_TAC, which will now be used by MESON instead of a bespoke variant. Added several new functions: SELECT_ELIM_CONV, SELECT_ELIM_ICONV, SELECT_ELIM_TAC, LAMBDA_ELIM_CONV and CONDS_ELIM_CONV. Sat 2nd Oct 04 trivia.ml, ind-types.ml Renamed the type "one" to "1". Didn't change any theorem names, but also added the obvious one_INDUCT and one_RECURSION and included them in the inductive type store (they were forgotten before). Also defined types "2" (with 2 elements) and "3" (with 3 elements), since I think all of them will be natural when using low-dimensional vectors ("real^3" etc.) Fri 1st Oct 04 sets.ml Changed FINITE_INTER to |- !s t. FINITE s \/ FINITE t ==> FINITE (s INTER t) where before it had a conjunction in the assumption, gratuitously weak. Fri 1st Oct 04 bool.ml Changed the bound variable names in the connective definitions to be more intuitive, e.g. "p" (not "P" or "t1") for Booleans, as suggested by Heath Putnam. Thu 30th Sep 04 term.ml, thm.ml Completed the basic "LCF-ization" by making terms and theorems into proper abstract types. Thu 30th Sep 04 drule.ml Removed use of "paconv" in PART_MATCH; I wanted to hide "paconv" and the new version is more efficient anyway. Wed 29th Sep 04 type.ml Made "hol_type" an abstract type, the first step in the more rigorous "LCF-ification". This was fairly painless; the only hidden value is a "the_type_constants" reference, which seems not to be used anywhere anyway. Wed 22nd Sep 04 hol.ml Modified the default setting of "hol_dir" so that if the "HOLDIR" environment variable isn't set, it will default to the current directory. The problem with the existing default was pointed out by Joe Hurd: it only really fits with my personal setup, whereas the current directory fits the intended build process. Wed 23rd Jun 04 sets.ml Added IMAGE_EQ_EMPTY = |- !f s. (IMAGE f s = {}) = (s = {}) and FORALL_IN_IMAGE = |- (!y. y IN IMAGE f s ==> P y) = (!x. x IN s ==> P(f x)) Mon 7th Jun 04 simp.ml Fixed an error in the treatment of "higher-order" congruence rules. Still, these need to be chosen with care otherwise beta-redexes arise in matching and send us into a loop. Wed 2nd Jun 04 parser.ml, printer.ml, sets.ml Added a new type of finite Cartesian products with a supported infix syntax `:A^B`, an indexing function so elements of `x` can be accessed by `x$1` ... `x$n`, and even a syntax for lambdas. If the type A is infinite, then it forces a 1-element type, otherwise uses an isomorphic image of A. The immediate motivation for this was multivariate calculus, but given that parser and printer support is needed anyway it seemed worth adding to the core. Various theorems have been included at the end of "sets.ml". Wed 2nd Jun 04 thm.ml Modified the basic type definition function so it sorts the type variables into alphabetical order. It's a pity to put sorting into the core, but we could if desired use some ultra-simple version. Without this there's no way of enforcing the right order of type variables, in particular for Cartesian products being defined. Wed 2nd Jun 04 sets.ml Added HAS_SIZE_IMAGE_INJ = |- !f s n. (!x y. x IN s /\ y IN s /\ (f(x) = f(y)) ==> (x = y)) /\ s HAS_SIZE n ==> (IMAGE f s) HAS_SIZE n Wed 12th May 04 sets.ml Strengthened num_FINITE from an implication to an equality: |- !s:num->bool. FINITE s = ?a. !x. x IN s ==> x <= a. Thu 6th May 04 preterm.ml Completely rewrote the typechecker. The initial motivation was to fix the fact that overload resolution was not working when the possible types could contain type variables, e.g. "N->real" for addition in multivariate calculus. But I took the chance to completely rewrite the typechecker in a more functional style so that I understand it, expunging the last non-trivial piece of hol90 code remaining in HOL Light. I suspect the new version is less efficient, but this is not really an issue for typical use. Sat 24th Apr 04 hol.ml Modified the "use_file" function used by "loads" and "loadt" so that it just returns quietly in the event of failure instead of raising an exception. The former behaviour resulted in OCaml rolling back the symbol table to the point before the load. The new behaviour is compatible with "#use" itself. However, I'd really like to make everything stop on the first error... Fri 23rd Apr 04 pa_j.ml, Makefile, .ocamlinit Replaced the old pa_j.ml, which doesn't work under 3.07, with a modified version kindly provided by Carl Witty. Also added a Makefile encapsulating the simple build process he pointed out to me. Fri 16th Apr 04 normalizer.ml Added parametrization over a variable ordering instead of assuming the default ordering. The impetus was the desire to re-use this inside the Cooper quantifier elimination procedure, in which case we want to use the order of quantifiers to decide the order of variables for easiest elimination. Also added SEMIRING_NORMALIZERS_CONV which exposes a whole suite of arithmetic operation conversions as well as the overall normalizer. Thu 8th Apr 04 lib.ml Renamed "funset" (map a finite partial function to a set representation of its graph) to just "graph". Wed 31st Mar 04 preterm.ml Fixed a bug in overloading, which was not following its promise to utilize polymorphic matching rather than simple equality among the possible alternative types. It was always assigning the generic type instead of the one resulting from matching, which of course results in an overall typechecking failure. Wed 31st Mar 04 arith.ml Added SUC_SUB1 = |- !n. SUC n - 1 = n, which I often use but is not in the main core arith file. Tue 30th Mar 04 theorems.ml Added the following two theorems, which I keep on regenerating for proofs: IMP_CONJ = |- p /\ q ==> r = p ==> q ==> r IMP_IMP = |- p ==> q ==> r = p /\ q ==> r Thu 26th Feb 04 real.ml Added "without loss of generality assume x <= y" (or "x < y") lemmas: REAL_WLOG_LE = |- (!x y. P x y = P y x) /\ (!x y. x <= y ==> P x y) ==> (!x y. P x y) REAL_WLOG_LT = |- (!x. P x x) /\ (!x y. P x y = P y x) /\ (!x y. x < y ==> P x y) ==> (!x y. P x y) Tue 3rd Feb 04 arith.ml Added a couple of theorems that embody some "exclusion zone" reasoning about integer quotients, and are quite tedious to derive for the cases in hand. DIV_LE_EXCLUSION = |- !a b c d. ~(b = 0) /\ b * c < (a + 1) * d ==> c DIV d <= a DIV b DIV_EQ_EXCLUSION = |- b * c < (a + 1) * d /\ a * d < (c + 1) * b ==> (a DIV b = c DIV d) Tue 25th Nov 03 wf.ml Added a more general form of the wellfounded recursion theorem with an "inductive invariant", based on the paper by Krstic and Matthews in TPHOLs'03. As they point out, this is much more useful for nested recursions like: `?g. !x. g(x) = if x = 0 then 0 else g(g(x - 1))` The proof can be condensed a bit by relying on MESON more but it gets rather slow in that case: let WF_REC_INVARIANT = prove (`WF(<<) ==> !H S. (!f g x. (!z. z << x ==> (f z = g z) /\ S z (f z)) ==> (H f x = H g x) /\ S x (H f x)) ==> ?f:A->B. !x. (f x = H f x)`, let lemma = prove_inductive_relations_exist `!f:A->B x. (!z. z << x ==> R z (f z)) ==> R x (H f x)` in REWRITE_TAC[WF_IND] THEN REPEAT STRIP_TAC THEN X_CHOOSE_THEN `R:A->B->bool` STRIP_ASSUME_TAC lemma THEN SUBGOAL_THEN `!x:A. ?!y:B. R x y` (fun th -> ASM_MESON_TAC[th]) THEN FIRST_X_ASSUM MATCH_MP_TAC THEN SUBGOAL_THEN `!x:A y:B. R x y ==> S x y` MP_TAC THEN ASM_MESON_TAC[]);; Mon 1st Sep 03 term.ml Fix for a second soundness bug, this time found by Bob Solovay; it was later independently discovered by Kevin Watkins. There was an asymmetry in the way alpha-conversion made its tests, which didn't check the proper correspondences between binding instances. let A = `\z:num. \z:num.z` and B = `\x:num. \y:num. x` in CONV_RULE NUM_REDUCE_CONV (BETA_RULE (AP_THM (AP_THM (ALPHA B A) `1`) `2`));; The new version, instead of using "assoc", traverses the environment more carefully checking proper correspondence. Wed 27th Aug 03 lib.ml Added "ran" (range of a finite partial function). Mon 18th Aug 03 normalizer.ml Added a new file, with a generic (semi)ring normalizer and its instantiation to the natural numbers. Later, this will be used more, e.g. in REAL_ARITH instead of the current mediocre version. Fri 15th Aug 03 sets.ml Added operations on set enumerations: dest_setenum, is_setenum, mk_setenum, mk_fset and some conversions SETIFY_CONV and SET_UNION_CONV. I suppose I should add a whole set eventually. Tue 22nd Jul 03 term.ml Did a bit of tidying up of the revised "inst" function, reinstating Boultonization of the Abs case and removing a redundant "qtry". Actually, we could now remove "qtry" from "lib.ml", and indeed get rid of "qcomb[2]" with only one change to the core. Fri 19th Jul 03 term.ml Made a preliminary fix to "inst" in order to fix a serious error leading to a soundness bug found by Tom Hales (the first soundness bug since 1996). The "unchanged" short-circuiting was bypassing the test for consistency with the environment, which can still fail even when the current variable is unchanged. I still need to think more carefully about this code, though. Here is my simplification of Tom's way of reaching a false theorem: let th0 = prove (`(?p:bool#B. P p) ==> ?x y. P(x,y)`, REWRITE_TAC[EXISTS_PAIR_THM]);; let th1 = CONV_RULE (RAND_CONV (BINDER_CONV (GEN_ALPHA_CONV `x:B`))) th0;; let th2 = INST_TYPE [`:bool`,`:B`] th1;; Mon 5th May 03 lib.ml and passim Made the trivial change of changing "sort" from uncurried to curried (the expected order is now 'a -> 'a -> bool not 'a # 'a -> bool). This meant a lot of trivial changes elsewhere, but it's worth getting it right once and for all. I also added all the "finite partial" functions stuff, which I've found very useful in my book code. This is useful for implementing some lookup structures a bit more efficiently. This also requires some changes elsewhere, changing "apply" to "apply_prover" in "simp.ml". Fri 4th Apr 03 arith.ml Added EVEN_ODD_DECOMPOSITION = |- !n. (?k m. ODD m /\ (n = 2 EXP k * m)) = ~(n = 0) Fri 7th Mar 03 list.ml Added APPEND_EQ_NIL = |- !l m. (APPEND l m = []) = (l = []) /\ (m = []) Thu 6th Mar 03 wf.ml Added a tactical and tactic WF_INDUCT_THEN and WF_INDUCT_TAC to perform wellfounded induction over a natural number measure. I do this quite often and I'm getting fed up with manually introducing a tweaked goal with a new variable etc. Wed 5th Mar 03 sets.ml Added CARD_IMAGE_LE = |- !f s. FINITE s ==> CARD(IMAGE f s) <= CARD s Tue 4th Mar 03 sets.ml Added a "dependent" generalization of a previous theorem: FINITE_PRODUCT_DEPENDENT = |- !s t. FINITE s /\ (!x. x IN s ==> FINITE (t x)) ==> FINITE {x,y | x IN s /\ y IN t x} Thu 27th Feb 03 parser.ml Made carriage return ("\r", 0x0d) another space. This is helpful on Windows, which normally uses CR/LF pairs for newline. Mon 24th Feb 03 sets.ml Added two theorems about cardinality of Cartesian products: CARD_PRODUCT = |- !s t. FINITE s /\ FINITE t ==> (CARD {x,y | x IN s /\ y IN t} = CARD s * CARD t) HAS_SIZE_PRODUCT = |- !s m t n. s HAS_SIZE m /\ t HAS_SIZE n ==> {x,y | x IN s /\ y IN t} HAS_SIZE m * n as well as "<=" versions of existing theorems for "<": HAS_SIZE_NUMSEG_LE = |- !n. {m | m <= n} HAS_SIZE n + 1 CARD_NUMSEG_LE = |- !n. CARD {m | m <= n} = n + 1 FINITE_NUMSEG_LE = |- !n. FINITE {m | m <= n} Fri 31st Jan 03 int.ml Removed spurious "integer" theorems that keep in the type destructors: INT_INV_0, INT_MUL_LINV, INT_POW_INV. Also tweaked INT_OF_REAL_THM so it handles conditionals (by adding a theorem to push "dest_int" through a conditional), and thus INT_POW_NEG is now what it was always supposed to be instead of having the spurious "dest_int"s. Fri 31st Jan 03 int.ml Added more pseudo-definitions (they are definitions on the reals) INT_GT, INT_LT and INT_SUB. Tue 28th Jan 03 int.ml Added INT_LT_TOTAL, which had got left out (probably because it was one of the real-closed field "axioms" over the reals). Tue 26th Nov 02 passim Made various changes to make HOL Light work in OCaml 3.06. The Camlp4 preprocessing can now be switched on and off with the magic identifiers "set_jrh_parsing" and "unset_jrh_parsing". Also set things up to look in $HOME/holl before defaulting to "johnh". Fri 1st Nov 02 real.ml Added another couple of handy theorems about finite sums: SUM_MORETERMS_EQ = |- !m n p. n <= p /\ (!r. m + n <= r /\ r < m + p ==> (f r = &0)) ==> (sum (m,p) f = sum (m,n) f) SUM_DIFFERENCES_EQ = |- !m n p. n <= p /\ (!r. m + n <= r /\ r < m + p ==> (f r = g r)) ==> (sum (m,p) f - sum (m,n) f = sum (m,p) g - sum (m,n) g) Fri 1st Nov 02 arith.ml Another basic theorem about DIV and MOD added. Will it never end? DIV_MONO2 = |- !m n p. ~(p = 0) /\ p <= m ==> n DIV m <= n DIV p Thu 31st Oct 02 real.ml Added SUM_EQ_0 = |- (!r. m <= r /\ r < m + n ==> (f r = &0)) ==> (sum (m,n) f = &0) There were already several theorems of that kind, but none of them quite in the natural convenient form I keep wanting. Thu 17th Oct 02 real.ml Added SUM_SWAP = |- !f m1 n1 m2 n2. sum (m1,n1) (\a. sum (m2,n2) (\b. f a b)) = sum (m2,n2) (\b. sum (m1,n1) (\a. f a b)) Wed 16th Oct 02 real.ml Added SUM_SPLIT = |- !f n p. sum(m,n) f + sum(m + n,p) f = sum(m,n + p) f, a useful generalization of SUM_TWO, which required a zero start. Tue 15th Oct 02 arith.ml Added DIV_MONO_LT = |- !m n p. ~(p = 0) /\ m + p <= n ==> m DIV p < n DIV p Mon 14th Oct 02 class.ml Finally added the following theorem, which I should use consistently instead of CONTRAPOS_CONV. I keep generating it by hand... let CONTRAPOS_THM = TAUT `!t1 t2. (~t1 ==> ~t2) = (t2 ==> t1)`;; Tue 13th Aug 02 wf.ml Couldn't resist putting in the simpler proof of WF_REC, where by using FIRST_X_ASSUM, the MESON proof search is quite feasible. I'd used this as an example at my PaPS invited talk and discussed it with Freek, so had started looking at the proof again. Tue 14th May 02 lib.ml Tidied up "lib.ml" in order to make things more consistent between CAML Light and OCaml versions. Also tweaked "partition" to do eq-optimization of either "yes" or "no" results. Mon 5th May 02 lib.ml Changed "forall", "exists" and "forall2" to treat list lazily; the previous implementations in terms of itlist did not have this property! This triviality improves runtimes by 3.5% Wed 4th Apr 02 lib.ml, type.ml, recursion.ml, preterm.ml, real.ml, term.ml Set up "assocd" and "rev_assocd" to return given defaults rather than failure; made use of these to replace a few of "try [rev_]assoc ... with Failure _ ->" idioms. Perhaps this should be done by more general functions that take a continuation, to allow us to scrub most of the others? Wed 4th Apr 02 lib.ml, type.ml, term.ml, meson.ml Changed "qmap" to propagate pointer equality not use the "Unchanged" exception. Modified type substitution and the meson FOL shadow syntax operations to work this way. Still need to do the same to terms. Tue 5th Mar 02 term.ml, basics.ml, thm.ml Moved a definition of "vfree_in" from basics.ml into term.ml and added a few function "freesin". Modified the inference rules to use these instead of "frees". This was motivated by the fact that GEN "x" (!x. Big[x] |- small[x]) took forever, since a full "frees" of the big assumption is performed. Admittedly this expands the size of the core, but the new functions are even simpler than "frees". It would perhaps be nice to eliminate "frees" completely from the core; it's now only used in variant-finding, and this could presumably be modified to use the new functions. Mon 4th Mar 02 pa_j.ml Added a new extension to bind the last toplevel expression to "it". The code was provided by Daniel de Rauglaudre in response to my question on the CAML list, following on from someone else wanting the same thing. Wed 13th Feb 02 drule.ml Further modified "new_definition" so that it uses recalled benign definitions correctly rather than taking the underlying definition as the overall theorem. Tue 12th Feb 02 drule.ml Modified "new_definition" so that it does in fact enter definitions in the store used to accept benign redefinitions --- previously this store was checked but not updated! Fri 8th Feb 02 sets.ml Changed the mis-named FINITE_RECUSION_DELETE to FINITE_RECURSION_DELETE. Fri 8th Feb 02 Ported HOL Light from CAML Light to OCaml (3.04), with hacked syntax to make "*" the special identifiers reserved for type constructors and modules, and with perverted-comma `...` quotations as well as <<...>>. The syntax modifier is kept in "pa_j.ml". Some minor syntactic changes were needed in most filed; only the following are unchanged: itab.ml, list.ml, nets.ml, num.ml, sets.ml, theorems.ml, wf.ml The names of useful functions and theorems are maintained, with the following exceptions (and lazy -> lazify but I never use that): Sum --> sum (the real number sum contravened even modified case conventions) assert -> check (assert is a reserved word with incompatible behaviour) Wed 6th Feb 02 arith.ml Added DIV_ADD_MOD = |- !a b n. ~(n = 0) ==> (((a + b) MOD n = a MOD n + b MOD n) = ((a + b) DIV n = a DIV n + b DIV n)) and DIV_REFL = |- !n. ~(n = 0) ==> (n DIV n = 1) and MOD_LE = |- !m n. ~(n = 0) ==> m MOD n <= m Tue 5th Feb 02 sets.ml Added ITSET_EQ = |- !s f g b. FINITE s /\ (!x. x IN s ==> (f x = g x)) /\ (!x y s. ~(x = y) ==> (f x (f y s) = f y (f x s))) /\ (!x y s. ~(x = y) ==> (g x (g y s) = g y (g x s))) ==> (ITSET f s b = ITSET g s b) Wed 30th Jan 02 int.ml Added INT_ABS_MUL_1 = |- !x y. (abs (x * y) = &1) = (abs x = &1) /\ (abs y = &1) Fri 25th Jan 02 int.ml Added INT_ABS_NUM and INT_ABS_MUL Wed 23rd Jan 02 sets.ml Added FINITE_DIFF = |- !s t. FINITE s ==> FINITE (s DIFF t) Tue 22nd Jan 02 int.ml Added INT_LE_SQUARE, another forgotten theorem. Wed 9th Jan 02 lib.ml, gtt.ml/hol.ml, caml, hol Stimulated by a message from Robert Solovay, re-organized and simplified the way the system is set up and loaded. Removed all use of the "Unix" library, so that one doesn't need a special CAML Light toplevel but can just use "camllight camlnum". (The use of "sys__command" still assumes a Unix-like environment, but this could also be changed quite easily.) The changes were: * Removing the "Interrupt" exception and the signal setup for it, just using CAML Light's native "Break", making sure "catch_break true" is set. * Using sys__time instead of unix__times in the "time" function * Avoiding using the PID to create unique temporary filenames, instead repeatedly adding "_" to the last component of the filename until it does not coincide with an existing file. (The file manipulation stuff in "sys" is a bit limited, so the code is messy.) Modified the setup to allow more flexibility over the placement of the filter file and avoid forcing the user to work in the HOL Light directory itself. The HOL directory and filter location are now specified by macros HOLDIR and HOLFILTER in the "hol" and "caml" scripts, and the load sequence attempts to read these if possible. These use new functions "includes" and "loads" so that HOL can be run from any location, and the original function "loadt" now uses a settable load path "load_path" which by default is the current directory and the HOL system directory. Wed 21st Nov 01 int.ml Added a whole new raft of theorems that were forgotten, to do with integer powers. INT_ABS_POW, INT_LE_POW2, INT_LT_POW2, INT_POW_1, INT_POW_1_LE, INT_POW2_ABS, INT_POW_2, INT_POW_ADD, INT_POW_EQ_0, INT_POW_INV, INT_POW_LE_1, INT_POW_LE2, INT_POW_LE, INT_POW_LT2, INT_POW_LT, INT_POW_MONO, INT_POW_MONO_LT, INT_POW_MUL, INT_POW_NEG, INT_POW_NZ, INT_POW_ONE, INT_POW_POW Thu 15th Nov 01 int.ml Added INT_OF_NUM_SUB = |- !m n. m <= n ==> (&n - &m = &(n - m)) which had somehow got left out, despite the analogous result for reals. Thu 15th Nov 01 arith.ml Added LE_EXP, a natural counterpart to LT_EXP. I'd avoided adding it until now (doubtless because it's so hideous with all the special cases). |- !x m n. x EXP m <= x EXP n = if x = 0 then (m = 0) ==> (n = 0) else (x = 1) \/ m <= n Wed 24th Oct 01 bool.ml Stimulated by a message from Mike Gordon on the hol-developers list, I discovered that the time taken for GEN on a theorem with an empty assumption list was not constant. This turned out to be because the MP with the proforma theorem was comparing (... = (\x. T)) and (... = (\v. T)), and this tips "aconv" into the full traversal instead of quick equality. Added an explicit alpha conversion to GEN so that this will now succeed instantly. Actually, several of these derived rules should avoid MP and use hypotheses, but that's another issue. Thu 5th Jul 01 preterm.ml Fixed a bug that Freek Wiedijk hit when trying to parse a set abstraction {t[x] | P[x]} with multiple instances of the same variable in the term t[x]. It turned out that "pfrees" was not returning a set, and could repeat the same variable because "ptm::acc" was used instead of "insert ptm acc", and this led to the problem since the variable then got existentially quantified twice, the outer one picking up a polymorphic type. Thu 28th Jun 01 basics.ml Fixed a bug in "type_match" pointed out by Michael Norrish; embarrassingly I thought I'd fixed this before (see 14 Feb 01). This was not recording identity mappings "A|->A" in the environment so far, meaning that it could allow impossible matches like `:A#A` to `:A#num`. Fixed it simply by removing the first line "if vty = cty then sofar else". Thu 21st Jun 01 sets.ml Added: CARD_IMAGE_INJ = |- !f s. (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y)) /\ FINITE s ==> (CARD(IMAGE f s) = CARD s) HAS_SIZE_POWERSET = |- !s n. s HAS_SIZE n ==> {t | t SUBSET s} HAS_SIZE 2 EXP n FINITE_POWERSET = |- !s. FINITE s ==> FINITE {t | t SUBSET s} IMAGE_DELETE_INJ = |- (!x. (f x = f a) ==> (x = a)) ==> (IMAGE f (s DELETE a) = IMAGE f s DELETE f a) Tue 22nd May 01 wf.ml In fact, it's even better; the reverse Skolemization is unnecessary! I took it out, and MESON can still do everything itself. In fact, with a few minutes, it can even absorb the shorter proof: REWRITE_TAC[WF_IND] THEN REPEAT STRIP_TAC THEN X_CHOOSE_THEN `R:A->B->bool` (ASSUME_TAC o last o CONJUNCTS) lemma THEN SUBGOAL_THEN `!x:A. ?!y:B. R x y` (fun th -> ASM_MESON_TAC[th]) THEN FIRST_ASSUM MATCH_MP_TAC THEN ASM_MESON_TAC[] Mon 21st May 01 wf.ml Simplified the proof of WF_REC using "reverse Skolemization" to hide the messy details of converting a relation into a function. This is now appealingly short, if tricky. Wed 9th May 01 int.ml Some more escapees: INT_ENTIRE, INT_EQ_MUL_LCANCEL, INT_EQ_MUL_RCANCEL, INT_LT_LMUL_EQ, INT_LT_RMUL_EQ. Wed 9th May 01 printer.ml Corrected a few irritating misfeatures with printing of numerals like "&1". First, avoided space after "&" for integers as well as reals. Second, made a space appear after "--" in "-- &1" etc. Really, we should have a more principled treatment of spacing based mainly on whether we end up with two adjacent printer tokens in the same lexical category (alphanumeric or symbolic). One day the printer will be rewritten. Sun 7th May 01 lib.ml Modified the "time" function so that it also reports elapsed times when a function generates an exception, i.e. fails or is interrupted. In addition, moved back the testing of "report_timing" to avoid doing anything special at all when it's false. Sat 6th May 01 int.ml Added INT_LT_MUL, another escapee. Fri 5th May 01 int.ml As part of writing a "proforma theorem" version of Michael Norrish's hol98 Cooper algorithm, I improved the integer theory a bit (this is the first time I've ever used it non-trivially). Added INT_LT_REFL, INT_LE_LMUL and INT_SUB_LDISTRIB, which had somehow got forgotten. Exposed INT_FORALL_POS at the top level (previously it was hidden inside ARITH_RULE, but is sometimes useful). Added the Archimedian theorem INT_ARCH = |- !x d. ~(d = & 0) ==> (?c. x < c * d) Fri 13th Apr 01 list.ml Added ALL2_ALL = |- !P l. ALL2 P l l = ALL (\x. P x x) l Mon 9th Apr 01 list.ml Added ALL2_MAP2 = |- !l m. ALL2 P (MAP f l) (MAP g m) = ALL2 (\x y. P (f x) (g y)) l m AND_ALL2 = |- !P Q l m. ALL2 P l m /\ ALL2 Q l m = ALL2 (\x y. P x y /\ Q x y) l m Sat 7th Apr 01 lib.ml Modified "forall2" to return falsity rather than fail when the two lists have different lengths. I had been under the impression it already worked this way. Fri 30th Mar 01 quot.ml Shortened the proof of the main proforma theorem by using MESON. I don't remember why I had such an intricate manual proof before. Actually, with better tweaks for extensionality and select terms, MESON should be capable of doing the whole thing automatically. Thu 29th Mar 01 thm.ml, bool.ml, drule.ml, class.ml Took the lists storing term and type variables out of the core "thm.ml". The discussion with Roger Jones on the Isabelle list reminded me of something Konrad Slind pointed out early on in the development of GTT: you don't need to store the actual definitions to ensure consistency. Renamed the basic definitional principle "new_basic_definition" by analogy with "new_basic_type_definition" and used this in bool.ml before the full definitional principle is defined, and also in class.ml to define "new_specification". However, re-introduced the list of definitions higher up in bool.ml, to enable acceptance of benign redefinitions. This was a nice feature that I used to have, but it got lost somewhere. This should be extended in two ways: allow appropriately named variables on the left (in case the definition was generated by code rather than entered via the quotation parser), and modify for higher-level derived rules like new_inductive_definition. And likewise for types. Sat 24th Mar 01 printer.ml Fixed a bug in the printing of "{}", which would print that whenever the head operator is the empty set, even if applied to an argument. This came up porting Peter Homeier's Church-Rosser proof. Wed 21st Mar 01 sets.ml Added a theorem about indexing of finite sets. Surprising I'd always managed without this so far: HAS_SIZE_INDEX = |- !s n. s HAS_SIZE n ==> (?f. (!m. m < n ==> f m IN s) /\ (!x. x IN s ==> (?!m. m < n /\ (f m = x)))) Tue 6th Mar 01 tactics.ml Removed VALID from TAC_PROOF, at the suggestion of Freek Wiedijk. This is a rather inefficient way to do things since it leads to a double evaluation of the justification function. Besides, we might want to remove VALID from the system, since on non-empty lists of subgoals it calls mk_thm (though here it's only used on the empty list of subgoals). Instead, added a test and alpha-conversion coercion to "prove" itself, to ensure that the theorem returned is exactly the same as the original goal, even up to variable naming. Tue 27th Feb 01 sets.ml Added SUBSET_DIFF = |- !s t. (s DIFF t) SUBSET s Mon 26th Feb 01 sets.ml Added IMAGE_DIFF_INJ = |- (!x y. (f x = f y) ==> (x = y)) ==> (IMAGE f (s DIFF t) = (IMAGE f s) DIFF (IMAGE f t)) Sun 25th Feb 01 sets.ml Added IMAGE_SUBSET = |- !f s t. s SUBSET t ==> IMAGE f s SUBSET IMAGE f t Sun 25th Feb 01 list.ml Added MEM_EL = |- !l n. n < LENGTH l ==> MEM (EL n l) l Sun 25th Feb 01 sets.ml Added SET_OF_LIST_APPEND = |- !l1 l2. set_of_list (APPEND l1 l2) = set_of_list l1 UNION set_of_list l2 Sat 24th Feb 01 meson.ml Fixed a bug in MESON where an error trap was capturing all failures with Cut -> failwith "meson_expand" | _ -> including Interrupt, hence in some circumstances making it impossible to interrupt the search. Changed "_" to "Failure _". This has popped up occasionally over the years, but last time I looked I didn't notice it, probably because I was searching for "with _". Thu 22nd Feb 01 sets.ml Added FINITE_SUBSETS = |- !s. FINITE s ==> FINITE {t | t SUBSET s} Wed 21st Feb 01 sets.ml Added IN_SET_OF_LIST = |- !x l. x IN set_of_list l = MEM x l Wed 14th Feb 01 basics.ml Fixed a bug in "type_match". Strangely, it was failing to detect inconsistent matching assignments and allowing say [`:num`,`:A`; `:bool`,`:A`]. This must have been around for ages without causing (many) problems. Sun 11th Feb 01 sets.ml Added IMAGE_UNION = |- !f s t. IMAGE f (s UNION t) = IMAGE f s UNION IMAGE f t IMAGE_o = |- !f g s. IMAGE (f o g) s = IMAGE f (IMAGE g s) Sat 10th Feb 01 wf.ml Added a theorem asserting that tail-recursive recursion schemes are always satisfiable. WF_REC_TAIL = |- !P g h. ?f. !x. f x = (if P x then f (g x) else h x) This was pointed out to me by J Moore at the 2000 ACL2 workshop. It would have saved me a bit of tedious hacking around defining things like unification algorithms. Sat 10th Feb 01 num.ml Added num_CASES = |- !m. (m = 0) \/ (?n. m = SUC n) This is often quite useful, and was in HOL88. Thu 8th Feb 01 sets.ml Added: FINITE_PRODUCT = |- !s t. FINITE s /\ FINITE t ==> FINITE {x,y | x IN s /\ y IN t} CARD_DELETE = |- !x s. FINITE s ==> (CARD(s DELETE x) = (if x IN s then CARD s - 1 else CARD s)) Tue 6th Feb 01 sets.ml Added FINITE_SUBSET_IMAGE = |- !f s t. FINITE t /\ t SUBSET IMAGE f s ==> (?s'. FINITE s' /\ s' SUBSET s /\ t SUBSET IMAGE f s') The proof is a bit slow (15 seconds); maybe I should unpick more of the automation. I've already separated out the two subgoals, which improves things slightly. Anyway, Moore's Law will save me eventually. Thu 25th Jan 01 tactics.ml Updated "STRIP_ASSUME_TAC" to discard the theorem if it's alpha-convertible to an existing assumption. This seems sensible, and the update brings HOL Light in line with HOL88. The implementation is the same, based on DISCARD_TAC, but this is only introduced locally since it doesn't seem particularly useful. This incompatibility arose in porting the proofs concerning graph planarity by Yamamoto et al. to HOL Light from hol90. Tue 23rd Jan 01 sets.ml Uncommented proof (it used to blow up the space usage a long time ago) of IN_DELETE_EQ = |- !s x x'. (x IN s = x' IN s) = x IN s DELETE x' = x' IN s DELETE x and added the theorems: INTER_ACI = |- (p INTER q = q INTER p) /\ ((p INTER q) INTER r = p INTER q INTER r) /\ (p INTER q INTER r = q INTER p INTER r) /\ (p INTER p = p) /\ (p INTER p INTER q = p INTER q) UNION_ACI = |- (p UNION q = q UNION p) /\ ((p UNION q) UNION r = p UNION q UNION r) /\ (p UNION q UNION r = q UNION p UNION r) /\ (p UNION p = p) /\ (p UNION p UNION q = p UNION q) while renaming the following, inexplicably called INTER_ACI and never used: INSERT_AC = |- (x INSERT y INSERT s = y INSERT x INSERT s) /\ (x INSERT x INSERT s = x INSERT s) Thu 18th Jan 01 lib.ml, calc_num.ml, ind-types.ml, real.ml, tactics.ml Modified "upto" to be an infix and take two arguments. This seems a much nicer and more flexible arrangement, and I've been meaning to do it for ages. Of course, quite a few consequential changes were needed, mostly from "upto n" to "0 upto n". Fri 12th Jan 01 tactics.ml Made the assumption list numbering when a goal is printed work from top to bottom, i.e. give the oldest assumption number zero, even though it is actually the last element of the list. NB: this only affects printing, and has no effect on goal representation. This policy seems more logical and should certainly be better when using numbers to refer to assumptions (as Freek Wiedijk is thinking of doing), since an assumption number won't change unless a prior assumption is deleted. Even better might be to allocate default labels, which would then never change. Thu 11th Nov 00 printer.ml Made set enumerations like `{1,2}` print properly instead of as `1 INSERT 2 INSERT EMPTY`. This includes the case of EMPTY printing as {}. Noticed this defect while doing 4CT statement with Freek Wiedijk, so made this fix on the boat from Holland. Apparently this once worked a long time ago with a different printer-constant called ESPEC, which is no longer used. Wed 4th Oct 00 wf.ml Added WF_REFL = |- !x. WF (<<) ==> ~(x << x) Wed 4th Oct 00 list.ml Added: MAP_FST_ZIP = |- !l1 l2. (LENGTH l1 = LENGTH l2) ==> (MAP FST (ZIP l1 l2) = l1) MAP_SND_ZIP = |- !l1 l2. (LENGTH l1 = LENGTH l2) ==> (MAP SND (ZIP l1 l2) = l2) MEM_ASSOC = |- !l x. MEM (x,ASSOC x l) l = MEM x (MAP FST l) ALL_APPEND = |- !P l1 l2. ALL P (APPEND l1 l2) = ALL P l1 /\ ALL P l2 Note that MAP_FST_ZIP is actually true without the precondition, but this is an accident of the way ZIP is defined by recursion on the first argument. Wed 4th Oct 00 basics.ml Fixed a bug in "mk_gabs" which must have been there forever and I only noticed when neither PAIRED_BETA_CONV nor GEN_BETA_CONV worked for `(\(p1,p2). f(p1,p2)) (p1,p2)`. The problem was that "f" was also used for the internal variable, with a faulty variant procedure due to a type or thinko. The line: let f = variant (frees tm1 @ frees tm2) (mk_var("f",fty)) was originally let f = variant (frees tm1 @ frees tm1) (mk_var("f",fty)) Tue 3rd Oct 00 wf.ml Added WF_LEX_DEPENDENT = |- !R S. WF R /\ (!a. WF (S a)) ==> WF (\(r1,s1). \(r2,s2). R r1 r2 \/ (r1 = r2) /\ S r1 s1 s2) and now deduced WF_LEX as a special case. The extra generality makes almost no difference to the proof, so it seemed we might as well prove a stronger form. Mon 2nd Oct 00 list.ml Added EX_MEM = |- !P l. EX P l = (?x. P x /\ MEM x l) Fri 29th Sep 00 wf.ml Finished the proof of WF_LEX which was commented out (actually took it from an old multiset ordering proof). Also uncommented WF_POINTWISE. However, erased WF_TC since TC isn't defined at this point in the build, and there is a short proof in Examples/reduct.ml Fri 29th Sep 00 list.ml Defined ASSOC (as in the ML "assoc"), and ITLIST2, and ZIP. I'm suffering from CAML envy. Fri 29th Sep 00 trivia.ml Removed various unused stuff from this file, including the definitions and consequential theorems for ASSOC, COMM, FCOMM, RIGHT_ID, LEFT_ID and MONOID. They don't seem particularly general (I've never used them), and I happened to want the name ASSOC for association lists. Fri 29th Sep 00 list.ml Defined FILTER. Added theorems: FILTER_APPEND = |- !P l1 l2. FILTER P (APPEND l1 l2) = APPEND (FILTER P l1) (FILTER P l2) FILTER_MAP = |- !P f l. FILTER P (MAP f l) = MAP f (FILTER (P o f) l) MEM_FILTER = |- !P l x. MEM x (FILTER P l) = P x /\ MEM x l Thu 28th Sep 00 sets.ml Added FINITE_SET_OF_LIST = |- !l. FINITE (set_of_list l) EX_MAP = |- !P f l. EX P (MAP f l) = EX (P o f) l EXISTS_EX = |- !P l. (?x. EX (P x) l) = EX (\s. ?x. P x s) l FORALL_ALL = |- !P l. (!x. ALL (P x) l) = ALL (\s. !x. P x s) l MEM_APPEND = |-!x l1 l2. MEM x (APPEND l1 l2) = MEM x l1 \/ MEM x l2 MEM_MAP = |- !f y l. MEM y (MAP f l) = ?x. MEM x l /\ (y = f x) Mon 25th Sep 00 filter.c Added #include . Michael Beeson pointed out that this was missing. I'd just got used to the laxity of gcc's default, but it also complains given -Wall. Mon 25th Sep 00 ind-types.ml Fixed a bug that had been lurking undiscovered in recursive types for a long time. It was found by Kim Sunesen in hol98 by trying the following: define_type "repDatatype = repBool bool | repTuple (repDatatype list) | repMap (bool list) # (repDatatype list) | repSet (bool list) | repEnum (bool list)";; and Michael Norrish identified the fix. Instead of applying the 1-step denesting transformation just to any of the nested types, apply it to one that is not a proper subtype of another. Otherwise, "lift_type_bijections" may later need to cope with lifting isomorphisms through other free recursive types, which it isn't able to do. The fix just uses: let nty = hd (sort (fun (t1,t2) -> occurs_in t2 t1) rectys) in instead of let nty = hd rectys in Wed 20th Sep 00 tactics.ml Modified UNDISCH_THEN so it tests for alpha-convertibility rather than equality. UNDISCH_TAC already did this, but it was forgotten here. Wed 20th Sep 00 sets.ml Added FINITE_IMAGE_INJ_GENERAL = |- !f A s. (!x y. x IN s /\ y IN s /\ (f x = f y) ==> (x = y)) /\ FINITE A ==> FINITE {x | x IN s /\ f x IN A} HAS_SIZE_NUMSEG_LT = |- !n. {m | m < n} HAS_SIZE n CARD_NUMSEG_LT = |- !n. CARD {m | m < n} = n FINITE_NUMSEG_LT = |- !n. FINITE {m | m < n} INFINITE_NONEMPTY= |- !s. INFINITE s ==> ~(s = EMPTY) INFINITE_DIFF_FINITE = |- !s t. INFINITE s /\ FINITE t ==> INFINITE (s DIFF t) Thu 7th Sep 00 list.ml Added LENGTH_REPLICATE = !n x. LENGTH (REPLICATE n x) = n Sat 2nd Sep 00 parser.ml Relaxed the lexical rules to allow identifiers to be built up from a mixture of alphanumerics and symbolics provided adjacent strings of opposite types are connected by "_". This is a bit of a hack but allows natural idioms like "++_2" as identifiers. It's also pretty sure to be upwards compatible. Fri 25th Aug 00 sets.ml Added INFINITE_IMAGE_INJ = |- !f. (!x y. (f x = f y) ==> (x = y)) ==> !s. INFINITE s ==> INFINITE (IMAGE f s) Mon 21st Aug 00 printer.ml Made general binders print properly. Previously this only worked for binders with simple abstractions as bodies, but not pairs etc. This should now work. Tue 18th Apr 00 tactics.ml Added "ANTS_TAC", a slightly more refined version of an idiom I use a great deal. I had the habit of using the cruder W(C SUBGOAL_THEN (fun th -> REWRITE_TAC[th]) o funpow 2 lhand o snd) but this is better and has a short name. Simply splits off antecedent of antecedent as a separate subgoal. The tactic has a laborious tautology definition; with a bit of reordering, we could use ITAUT or even TAUT. Fri 24th Mar 00 list.ml Added ALL_MEM = |- !P l. (!x. MEM x l ==> P x) = ALL P l Tue 18th Jan 00 arith.ml Added MOD_ADD_MOD = |- !a b n. ~(n = 0) ==> ((a MOD n + b MOD n) MOD n = (a + b) MOD n) Tue 21st Dec 99 printer.ml Changed the "dest_binder" used by printer stuff to "dest_binder_vorc" which will break binders even when the binder thing is a variable. Otherwise the printer actually crashes when used on a binder that isn't a constant, as I discovered by trying to demonstrate how to define binders to Jim Grundy. Thu 16th Dec 99 arith.ml Added MOD_MULT2: |- !m n p. ~(m * p = 0) ==> ((m * n) MOD (m * p) = m * n MOD p) this was wanted by Jim Grundy, and in any case is a logical addition since the corresponding theorem for DIV (DIV_MULT2) was already there. Tue 14th Dec 99 calc_rat.ml Modified REAL_RAT_POW_CONV so that it works when the argument is a decimal. Previously for example this would fail: REAL_RAT_POW_CONV `#0.7854 pow 2` Thu 9th Dec 99 class.ml Modified COND_CASES_TAC so that if the expression being split over is a negation, the assumption in the second branch eliminates the double negation. This was a suggestion from Jim Grundy. Fri 3rd Dec 99 preterm.ml Fixed a bug in overloading found by Jim Grundy. It was the case that if there was enough type information to resolve an overload unambiguously, then this already-gathered type information was used as the type for the object, instead of taking the type of the overload instance from the interface list. However, this is no good when we have enough type information to resolve the overload but still have indeterminate parts, e.g. a binary function where we have the type of the first argument (and hence resolve the overload) but not the second. Not only doesn't the subsequent typecheck ever resolve the indeterminate parts, they actually look to it like incompatibilities, so it fails. This has been modified to simply take the type from the interface list, as is done anyway in the case of defaulting. Also, typechecking is now repeated whenever "overresolve" changed the preterm, not just when it set the "overloads_defaulted" flag. Fri 3rd Dec 99 printer.ml Put in a flag "reverse_interface_mapping" that controls whether overloaded identifiers print as the overloaded name (flag = true, default) or their true name (flag = false). This was a suggestion from Jim Grundy. Fri 3rd Dec 99 preterm.ml Backed off the change made a year ago that made "\const. tm" always interpret "const" as a variable. Now this behaviour is under the control of a settable flag called "ignore_constant_varstructs". However, the default is true, which means no (intentional) change in behaviour. Setting it to false would restore the state of affairs in version 1.0 Fri 3rd Dec 99 pair.ml Added the GEN_BETA_CONV function that does generalized beta conversion over nested linear patterns of constructors, e.g. GEN_BETA_CONV `(\[SUC m; SUC n]. m + n) [SUC 1; SUC 2]`. This subsumes PAIRED_BETA_CONV, but it is slower. Fri 3rd Dec 99 recursion.ml Added "create_projections" which demonstrates the existence of projection functions for a type constructor. These are then stored away in a cache. The main intention is to do genrealized betas over arbitrarily nested type constructors automatically. Fri 3rd Dec 99 drule.ml Added BETAS_CONV where BETAS_CONV n `(\x1...xn. E[xs]) a1 ... an` reduces all the beta redexes. Thu 2nd Dec 99 printer.ml Stuck an additional "print_break" before the "and"s in a printing of "let ... and ... and ... and", otherwise the line would never break properly. This bug was pointed out by Jim Grundy. Thu 2nd Dec 99 ind-type.ml Added ":num" to the inductive type store. I needed to swap the order of arguments; I should really make this standard anyway. Wed 1st Dec 99 arith.ml Added MOD_MULT_ADD = |- !m n p. (m * n + p) MOD n = p MOD n Mon 29th Nov 99 ind-types.ml Changed the recursive type definition type bijections used internally from "mk_" and "dest_" to "_mk_" and "_dest_". Jim Grundy was caught out by the original, since he wanted a constructor with the same name. Sun 22nd Aug 99 real.ml Added REAL_POW2_ABS = |- !x. abs(x) pow 2 = x pow 2 REAL_LE_SQUARE_ABS = |- !x y. abs(x) <= abs(y) = x pow 2 <= y pow 2 Sun 22nd Aug 99 ind-defs.ml Renamed MONO_ALL to MONO_FORALL; it's more logical and the second is already used for list quantification. Mon 3rd May 99 realax.ml, calc_rat.ml, calc_real.ml Set up "--" as a prefix, so one can write "--R(j)" rather than "--(R(j))" etc. This meant changing a few `--`s to `(--)`s otherwise the parser complains. Fri 23rd Apr 99 meson.ml Wow, found and removed a second "with Cut -> raise Cut". What fascinating optimizations there are to be found! Fri 9th Apr 99 meson.ml While nosing around meson.ml looking for Wishnu's bug (see last item) I spotted and removed a pointless "with Cut -> raise Cut". Might speed things up by a few microseconds. Fri 9th Apr 99 canon.ml Fixed a bug in NNFC_CONV. The "clever" optimization that avoided rewriting a term twice at one level is wrong in the case "~ ~ (p ==> q)". This caused a MESON_TAC failure spotted by Wishnu Prasetya: `~(p /\ q ==> ~r) = (p /\ q /\ r)`;; The new version simply uses TOP_SWEEP_CONV. I don't believe the impact will be that significant. These things aren't that well optimized anyway; if speed is critical I should rewrite it to do the maching manually. Fri 2nd Apr 99 tactics.ml Fixed a bug in THENL where "g THENL [non-empty-list]" would fail if g left no subgoals. Now it behaves like THEN. Thu 25th Mar 99 calc_num.ml Installed a more efficient version of NUM_ADD_CONV and NUM_SUC_CONV. These are hand-coded instantiations rather than being based on rewriting. The motivation is that addition is now fairly important in proofs, though the long-term goal is to improve multiplication, which is the real bottleneck. Anyway, this improved the speed of my current proofs (division algorithms) by about 12%. Wed 24th Mar 99 real.ml Added: REAL_EQ_RDIV_EQ = |- !x y z. &0 < z ==> ((x = y / z) = x * z = y) REAL_EQ_LDIV_EQ = |- !x y z. &0 < z ==> ((x / z = y) = x = y * z) Thu 18th Feb 99 sets.ml Renamed the old FINITE_IMAGE to FINITE_IMAGE_EXPAND and added: FINITE_IMAGE = |- !f s. FINITE s ==> FINITE (IMAGE f s) The previous version had the definition of IMAGE expanded... Thu 18th Feb 99 sets.ml Added IMAGE_CLAUSES = |- (IMAGE f EMPTY = EMPTY) /\ (IMAGE f (x INSERT s) = f x INSERT IMAGE f s) Thu 18th Feb 99 arith.ml Added: MOD_MULT_RMOD = |- !m n p. ~(n = 0) ==> ((m * p MOD n) MOD n = (m * p) MOD n) MOD_MULT_LMOD = |- !m n p. ~(n = 0) ==> ((m MOD n * p) MOD n = (m * p) MOD n) MOD_MULT_MOD2 = |- !m n p. ~(n = 0) ==> ((m MOD n * p MOD n) MOD n = (m * p) MOD n) MOD_EXP_MOD = |- !m n p. ~(n = 0) ==> ((m MOD n) EXP p MOD n = m EXP p MOD n) Wed 27th Jan 99 real.ml Added: REAL_LT_POW2 = |- !n. &0 < &2 pow n REAL_LE_POW2 = |- !n. &1 <= &2 pow n Wed 27th Jan 99 pair.ml Added EXISTS_PAIR_THM = |- (?p. P p) = (?p1 p2. P (p1,p2)) Fri 22nd Jan 99 real.ml Moved the quantifier from ... ==> !n. to !n. ... ==> in SUM_POS_GEN, so that we can later use MATCH_MP_TAC more conveniently. Wed 20th Jan 99 drule.ml Modified HIGHER_REWRITE_CONV so that it takes a flag allowing one to pick the lowest subterm (as originally) or the highest. Now the latter is used for COND_CASES_TAC, which seems much more sensible. However, note that lowest still makes sense for many things, e.g. the elimination of subtractions. And the new COND_CASES_TAC has the "defect" that if (if ...) will put the inner conditional on the assumptions so REPEAT COND_CASES_TAC is no good. Wed 20th Jan 99 pair.ml Modified let_CONV so that it deals with paired lets correctly provided the thing being abbreviated is also paired. Modified LET_TAC so that it works on any paired let at all, introducing appropriate abbreviatory assumptions. Wed 20th Jan 99 basics.ml, equal.ml Added three new functions "find_path", "follow_path" and "PATH_CONV" to create and use paths, i.e. director strings. These are useful for modifying terms at a precise position without any accidents. Thu 7th Jan 99 arith.ml Made "mk_numeral" and "mk_small_numeral" fail when given negative arguments. Thu 10th Dec 98 preterm.ml Modified the behaviour of "typify" on the varstruct of a preterm so that if the varstruct is a single variable, it is always treated as a variable even if there is a constant with that name. This makes (simply) bound variables override constants, conforming to the intuition that constants are like outer-level bound variables. However, this does not apply to complex varstructs; this seems hard in general since we need to recognize certain constants like the comma in "\(x,y). ...". Probably need to think more about this in concert with set abstraction syntax. Tue 8th Dec 98 arith.ml Added LE_SQUARE_REFL = |- !n. n <= n * n Fri 4th Dec 98 real.ml Added REAL_DIV_REFL = |- !x. ~(x = &0) ==> (x / x = &1) REAL_DIV_LMUL = |- !x y. ~(y = &0) ==> (y * (x / y) = x) REAL_DIV_RMUL = |- !x y. ~(y = &0) ==> ((x / y) * y = x) Thu 3rd Dec 98 arith.ml Added the new theorems: LE_LDIV = |- !a b n. ~(a = 0) /\ b <= a * n ==> b DIV a <= n LE_RDIV_EQ = |- !a b n. ~(a = 0) ==> (n <= b DIV a = a * n <= b) LE_LDIV_EQ = |- !a b n. ~(a = 0) ==> (b DIV a <= n = b < a * (n + 1)) Wed 2nd Dec 98 class.ml Modified the condition syntax of COND_CONG to the new "if g then t else e" instead of the old "g => t | e". Thu 12th Nov 98 arith.ml Added MOD_EQ_0 = |- !m n. ~(n = 0) ==> ((m MOD n = 0) = (?q. m = q * n)) EVEN_MOD = |- !n. EVEN(n) = (n MOD 2 = 0) ODD_MOD = |- !n. ODD(n) = (n MOD 2 = 1) Wed 11th Nov 98 arith.ml Added DIV_MONO = |- !m n p. ~(p = 0) /\ m <= n ==> m DIV p <= n DIV p DIV_EQ_0 = |- !m n. ~(n = 0) ==> ((m DIV n = 0) = m < n) Fri 23rd Oct 98 gtt.ml, hol.ml Set up an allocated scratch_directory that is used instead of /tmp when the environment variable TMPDIR isn't set. This was because, at least on my laptop, /tmp was getting cleared so often that it sometimes happened in the middle of loading a multiple-day proof. Tue 20th Oct 98 list.ml Changed the names of FORALL and FORALL2 to ALL and ALL2. The new names are shorter, more consistent with EX, and make it easier to name theorems relating ordinary quantifiers to these. For example the theorem previously known as FORALL_FORALL gives no clue in its name about which way it switches the quantifiers. The names of theorems have been changed as follows: FORALL -> ALL FORALL2 -> ALL2 FORALL_IMP -> ALL_IMP NOT_FORALL -> NOT_ALL FORALL_MAP -> ALL_MAP FORALL_T -> ALL_T FORALL_MP -> ALL_MP FORALL_FORALL -> FORALL_ALL AND_FORALL -> AND_ALL MONO_FORALL -> MONO_ALL FORALL2_DEF -> ALL2_DEF MAP_EQ_FORALL2 -> MAP_EQ_ALL2 FORALL2_MAP -> ALL2_MAP FORALL2_AND_RIGHT -> ALL2_AND_RIGHT MONO_FORALL2 -> MONO_ALL2 Tue 20th Oct 98 printer.ml Made sure that a space is printed between a prefix operator and its arguments, and between a binder and its bound variables, when (and only when) the operator's name is alphanumeric. Mon 19th Oct 98 class.ml, ind-defs.ml, itab.ml real.ml tactics.ml Removed "PROVE" since there's now no compactor, and replaced it by "prove". Mon 19th Oct 98 calc_num.ml Added EXPAND_CASES_CONV, which I seem to use quite often to expand "!n. n < N ==> P n" into "P 0 /\ P 1 /\ ... /\ P (N-1)". Mon 19th Oct 98 calc_rat.ml Added more functions for dealing with rational number terms: numdom, numerator, denominator, term_of_rat, rat_of_term. Fri 16th Oct 98 basics.ml, bool.ml, tactics.ml Eliminated the compactor and loading and saving of theorems. This was a hack and in fact turned out to inhibit performance on really big examples. Admittedly, this is probably just because I used the built-in CAML hash tables which clearly aren't designed for hashing objects of this size and structure. Fri 16th Oct 98 tactics.ml Optimized the implementations of THEN and THENL which were recreating nontrivial bits of the justification function; these are now going to be function closures. This saves piles of memory on really big case splits. Also made similar optimizations to a few other "terminal" tactics such as ACCEPT_TAC. Thu 15th Oct 98 real.ml, int.ml Renamed REAL_NEG_EQ_0 and INT_NEG_EQ_0 from REAL_NEG_EQ0 and INT_NEG_EQ0 for consistency with other REAL_opr_EQ_0. Really a more serious bash at making these names consistent would be worthwhile before they become entrenched. Thu 15th Oct 98 simp.ml Modified "net_of_thm" so it also translates |- p ==> (s = t[s]) into |- p ==> ((s = t[s]) = T). Previously this was only happening for non-conditional rewrites. Also eliminated the "discarding looping rewrite" warning which would need changing to "modifying looping rewrite" and is probably more irritating than useful. Fri 9th Oct 98 parser.ml Removed the special "bracket" status for semicolon; there seem to be few natural situations where a semicolon is followed by a symbolic ID. Anyway, I wanted to use ";;" for a sequencing operator. Maybe I should do the same for comma? Wed 7th Oct 98 tactics.ml Made goals print the other way up, i.e. with the conclusion at the bottom with the head (most recently added) assumptions above them. This is a trivial change, but seems much more sensible to avoid goals with zillions of assumptions scrolling off the screen. Tue 6th Oct 98 tactics.ml Reimplemented THEN and THENL in a more direct way, not relying on hacking up a goalstate and manipulating it using "by" and "bys". This is partly because I suspect the space behaviour of the old one is very bad when there are splits into many subgoals. Anyway the aesthetics are better. The left-right order for metavariable instantiation has been maintained. Tue 6th Oct 98 drule.ml Optimized INSTANTIATE_ALL for the case where the instantiation is null. Tue 6th Oct 98 tactics.ml Made similar optimization to the one below for THENL, this time for THEN, and hence EVERY. I had misthought that THEN called THENL, but of course it doesn't. Tue 6th Oct 98 real.ml Changed "prove" to "PROVE" so that REAL_ARITH doesn't cache what probably are, after all, only temporary theorems. Mon 5th Oct 98 real.ml Added call to "clear_atom_cache" in SIMPLE_REAL_ARITH_TAC; previously the cache would just extend forever across multiple calls, surely not a very good idea... Mon 5th Oct 98 tactics.ml Optimized the implementation of "THENL" for the case where the composite construct finishes off the goal. Now, rather than keeping the perhaps complicated justification function hanging around, we apply it at once, then return a trivial wrapper round the resulting theorem. Tue 15th Sep 98 real.ml Added REAL_DIV_POW2_ALT = |- !x m n. ~(x = &0) ==> (x pow m / x pow n = (if n < m then x pow (m - n) else inv (x pow (n - m)))) Mon 14th Sep 98 arith.ml Added MOD_EXISTS = |- !m n. (?q. m = n * q) = (if n = 0 then m = 0 else m MOD n = 0) Mon 14th Sep 98 pair.ml Made LET_TAC scrub trivial lets first. This is a bit ad-hoc, but helps it goals with parallel chains of lets but different bodies. Fri 4th Sep 98 arith.ml Added LT_MULT2 = |- !m n p q. m < n /\ p < q ==> m * p < n * q which had somehow got forgotten. Thu 3rd Sep 98 basics.ml, pair.ml Modified dest_let so that it allows arbitrary varstructs on the left of the equals sign. This was just a matter of changing "strip_abs" to "strip_gabs". Also modified let_CONV so that it works on paired varstructs provided the right hand side is a pair, as in "let (x,y,z) = (1,2,3) in ...". The parser already worked in the general case, and the change to "dest_let" makes the printer work. However let_CONV needs more work, changing ABS to an equivalent for varstructs. Pretty easy... Wed 2nd Sep 98 arith.ml Added EXP_1 = |- !n. n EXP 1 = n, and renamed the theorem previously called EXP_1 to EXP_ONE = |- !n. 1 EXP n = 1 Wed 2nd Sep 98 pair.ml Added LET_TAC to eliminate let_terms while replacing them with equivalently named abbreviatory assumptions. Sun 16th Aug 98 real.ml Added REAL_DIV_POW2 = |- !x m n. ~(x = &0) ==> (x pow m / x pow n = (if n <= m then x pow (m - n) else inv (x pow (n - m)))) Fri 14th Aug 98 calc_rat.ml Improved REAL_RAT_MUL_CONV which did cancellation in an utterly inept way, requiring big multiplications in many cases. Fri 14th Aug 98 real.ml Term-netted REAL_INT_RED_CONV and so REAL_INT_REDUCE_CONV. Fri 14th Aug 98 real.ml Added REAL_LT_DIV = |- !x y. &0 < x /\ &0 < y ==> &0 < x / y REAL_LE_DIV = |- !x y. &0 <= x /\ &0 <= y ==> &0 <= x / y Tue 4th Aug 98 calc_rat.ml The conversion CALC_DIV_CONV was failing if the input was of the form "m / n" for integers m and n even if m and n were cancellable. Now it cancels if possible. Wed 29th Jul 98 arith.ml Added MOD_MOD_REFL = |- !m n. ~(n = 0) ==> (m MOD n MOD n = m MOD n) Mon 27th Jul 98 tactics.ml Added RECALL_ACCEPT_TAC, intended for use by extremely time-consuming rules that take some argument(s) other than the goal term. One can use it to delay evaluation till after lookup in the theorem compactor. Fri 24th Jul 98 calc_rat.ml Term-netted REAL_RAT_RED_CONV and so REAL_RAT_REDUCE_CONV. Thu 23rd Jul 98 pair.ml Added FORALL_PAIR_THM = |- (!p. P p) = (!p1 p2. P (p1,p2)) I should probably get rid of GEN_PAIR_TAC one day. Wed 22nd Jul 98 real.ml Added REAL_LE_RDIV_EQ = |- !x y z. &0 < z ==> (x <= y / z = x * z <= y) REAL_LE_LDIV_EQ = |- !x y z. &0 < z ==> (x / z <= y = x <= y * z) REAL_LT_RDIV_EQ = |- !x y z. &0 < z ==> (x < y / z = x * z < y) REAL_LT_LDIV_EQ = |- !x y z. &0 < z ==> (x / z < y = x < y * z) REAL_LT_DIV2_EQ = |- !x y z. &0 < z ==> (x / z < y / z = x < y) REAL_LE_DIV2_EQ = |- !x y z. &0 < z ==> (x / z <= y / z = x <= y) REAL_EQ_LCANCEL_IMP = |- !x y z. ~(z = &0) /\ (z * x = z * y) ==> (x = y) REAL_EQ_RCANCEL_IMP = |- !x y z. ~(z = &0) /\ (x * z = y * z) ==> (x = y) Wed 15th Jul 98 real.ml Added SUM_CONST = |- !c n. Sum (0,n) (\m. c) = &n * c Mon 13th Jul 98 real.ml Added REAL_INV_1_LE = |- x. &0 < x /\ x <= &1 ==> &1 <= inv x Fri 10th Jul 98 calc_num.ml Added NUM_FACT_CONV into the set of conversions applied by NUM_RED_CONV and hence NUM_REDUCE_CONV. Fri 10th Jul 98 real.ml Added REAL_POW_POW = |- !x m n. x pow m pow n = x pow (m * n) Thu 9th Jul 98 arith.ml Added MOD_MOD = |- !m n p. ~(n * p = 0) ==> (m MOD (n * p) MOD n = m MOD n) and DIV_MULT2 = |- !m n p. ~(m * p = 0) ==> ((m * n) DIV (m * p) = n DIV p) and MOD_1 = |- !n. n MOD 1 = 0 Mon 5th Jul 98 real.ml Added REAL_POW_MONO_LT = |- !m n x. &1 < x /\ m < n ==> x pow m < x pow n Mon 5th Jul 98 tactics.ml Made the variant-choosing in CHOOSE_TAC more conservative, to avoid problems where the variable coincides with something in the hypotheses of an assumption theorem, yet not in any of the conclusions. Mon 5th Jul 98 real.ml Added another couple: REAL_LT_RMUL_EQ = |- &0 < z ==> (x * z < y * z = x < y) REAL_LT_LMUL_EQ = |- &0 < z ==> (z * x < z * y = x < y) Thu 2nd Jul 98 real.ml Added two theorems that are specially useful in conjunction with the simplifier: REAL_LE_RMUL_EQ = |- &0 < z ==> (x * z <= y * z = x <= y) REAL_LE_LMUL_EQ = |- &0 < z ==> (z * x <= z * y = x <= y) Thu 25th Jun 98 realax.ml Avoided the auxiliary definition of "sum". This was inelegant, and particularly inconvenient since "sum" is quite a useful constant name for the user. Wed 24th Jun 98 tactics.ml Added EXPAND_TAC, the natural complement to ABBREV_TAC. Tue 23rd Jun 98 parser.ml Added comments to the quotation parser. These are one-line comments started by a settable symbol, initially // as in BCPL/C++. Tue 26th May 98 calc_rat.ml, parser.ml, printer.ml Made `#xxxx.yyyy` acceptable instead of explicit fraction, using a new constant DECIMAL. Modified parser and printer appropriately and tweaked rational arithmetic stuff so it accepts these (but always returns a conventional rational). Wed 20th May 98 preterm.ml Made hexadecimal numerals acceptable to the parser. They are entered in C notation, i.e. 0x... with either upper or lower case or mixed case for the hex digits (but the initial x must be lowercase). Note that this has no logical meaning, so the numerals are still printed as decimals. However it would be possible to have a more general numeral syntax, e.g. several alternatives to "NUMERAL". This would mean changing some calculation code. ********************** RELEASE OF VERSION 1.00 ********************** Thu 30th Apr 98 preterm.ml, parser.ml, printer.ml, passim Abolished interface maps, integrating them with overloading. Renamed the functions involved, except make_overloadable. Now we have override_interface, overload_interface, reduce_interface and remove_interface. Wed 29th Apr 98 tactics.ml Deleted test_just. I think the stuff is all OK, so I shouldn't need it; and it cheats! Wed 29th Apr 98 arith.ml Fixed a bug in PRE_ELIM_TAC, used in preprocessing phase of the decision procedure for naturals. It was substituting x = SUC m where we had PRE(x), but of course x might not be a variable and then the x = SUC m was lost. The analogous SUB_ELIM_CONV already did this right. Wed 29th Apr 98 simp.ml, theorems.ml, class.ml Tweaked simplification slightly, adding new rewrites for "(x = x) ==> p" and "if x = x then y else z" so that we never perform a contextual rewrite in such a trivial case. Wed 29th Apr 98 meson.ml Limited the case-splitting before MESON proper is applied; this was getting to rather extreme levels for propositional tautologies. The limit is controlled by an assignable variable meson_split_limit; this is initially 4, meaning no more than 2^4 = 16 separate cases. Wed 29th Apr 98 real.ml Optimized the linear decision procedure in several ways. (1) the most important: the case-splitting induced by "abs" is now somewhat more intelligent that blindly doing COND_CASES as before. Something like abs(x) <= y is expanded to x <= y /\ --x <= y, which is much better: no case-splitting. In the case split from x <= abs(y) we have a case split but leave out &0 <= x and &0 <= -x indications; they are not needed and if "x" involves abs, can lead to further expansion. It's still not perfect; for example &2 * abs(x) <= y is still treated in the old way. But it's much better. (2) We avoid ASSUMEing and PROVE_HYPing the starting theorems. (3) simplified the proof traces slightly, which for no good reason kept the shadow forms of the intermediate inequalities. Mon 27th Apr 98 thm.ml, drule.ml, class.ml, bool.ml Abandoned the "clever" definitional principle and made the equational form primitive. Mon 27th Apr 98 type.ml, term.ml, drule.ml, tactics.ml Renamed type_subst to raw_type_subst and made type_subst a version that doesn't create an unchanged exception; better for external use. Tue 21st Apr 98 int.ml Added some theorems whose real counterparts were in realax.ml not real.ml and hence got left out, e.g. INT_ADD_SYM and INT_ADD_ASSOC (!) Tue 21st Apr 98 pair.ml Added let_CONV. Tue 31st Mar 98 type.ml, term.ml Made a few efficiency tweaks and tidyings. Tue 31st Mar 98 type.ml Removed "tyvarsl" which was never used. Wed 25th Mar 98 list.ml Added definition of constant EL. Tue 24th Mar 98 printer.ml Forced the printer to print brackets round paired abstractions. Mon 23rd Mar 98 printer.ml Stopped printing of CONS as (CONS). I've no idea why I did this. Mon 2nd Mar 98 simp.ml, class.ml Set up useful simplifier as the default SIMP_TAC, ASM_SIMP_TAC etc., for the situations where one needs implicational rewriting but doesn't want to fuss with constructing simpsets. Fitted mechanism for basic congruences and augmented it with conditional in class.ml Wed 18th Feb 98 printer.ml Made a modification to compile under 0.74: print_break is now curried. Wed 11th Feb 98 ind-types.ml Changed "BOT" to "BOTTOM", less likely to be wanted elsewhere when dealing with Boolean algebras. Tue 10th Feb 98 printer.ml Modified the printer to use bracketing instead of dollaring. Tue 10th Feb 98 quot.ml, real.ml Modified "define_quotient_type" so the user can specify the names of the abstraction and representation functions. Mon 9th Feb 98 parser.ml, printer.ml, passim Eliminated the use of dollaring to force non-special status. Now just use brackets: (+), (==>), (!) etc. This seems much more intuitive and also makes the lexical stuff simpler. Sun 8th Feb 98 printer.ml Added user-installable printers, via a trivial "try them first" interface. Wed 4th Feb 98 list.ml Added EX_IMP. Wed 28th Jan 98 list.ml Added NOT_FORALL, FORALL_FORALL and AND_FORALL. Mon 19th Jan 98 list.ml Added FORALL_MP. Wed 14th Jan 98 list.ml Added ITLIST_EXTRA. Fri 2nd Jan 98 simp.ml Inserted a warning when discarding a looping rewrite. Wed 24th Dec 97 drule.ml Made INSTANTIATE fail if anything gets instantiated in assumptions. The main reason is that on occasion ASM_REWRITE_TAC[] could instantiate the assumptions of goal-assumptions and lead to an invalid tactic, whereas it should fail. Mon 22nd Dec 97 sets.ml Added LENGTH_LIST_OF_SET and rejigged the definition to make sure "list_of_set" is repetition-free. Mon 22nd Dec 97 list.ml Added MAP2. Sat 20th Dec 97 sets.ml Added num_FINITE, num_FINITE_AVOID and num_INFINITE. Fri 19th Dec 97 list.ml Added FORALL2_AND_RIGHT. Thu 18th Dec 97 sets.ml Added FINITE_IMAGE_INJ. Thu 18th Dec 97 list.ml Added MAP_EQ_DEGEN. Thu 18th Dec 97 ind-types.ml Made "prove_constructors_injective" and "prove_constructors_distinct" return a single conjunctive theorem rather than a list of theorems. This strange decision has been an irritation for ages, so I decided to change it now before it requires any more tweaking of work. Thu 18th Dec 97 list.ml Added FORALL2, with theorems MAP_EQ_FORALL2 and FORALL2_MAP; added in monotonicity theorems for FORALL and FORALL2 for ind-defs package. Also tweaked a couple of proofs. Wed 17th Dec 97 list.ml Added FORALL_T. Sat 13th Dec 97 sets.ml Added UNIONS_INSERT, FINITE_UNIONS and FINITE_IMAGE. Thu 11th Dec 97 sets.ml Added IN_UNIONS, IN_INTERS, UNIONS_0, UNIONS_1, UNIONS_2. Tue 2nd Dec 97 equal.ml Added CACHE_CONV, which seems potentially quite useful. I should probably use it in the reals decision procedure in future; it uses a term net rather than a linear list, as now. However this version also uses a modicum of intelligence. Tue 2nd Dec 97 tactics.ml Changed THENL so it succeeds if the first tactic creates no subgoals, even if the list of tactics is of nonzero length. This was suggested by Michael Norrish (for hol90). Sat 29th Nov 97 preterm.ml Updated the earlier change so it also checks for numerals, not actually there as single constants. At this stage they are not written out in binary. Fri 28th Nov 97 preterm.ml Fixed up set abstractions so they don't generalize over constants in the left; the problem was that at the preterm stage these are still "variables", so needed to put a check in "pfrees" whether something is a constant name. This is safe: any inner bound variables aren't free anyway(!) Fri 28th Nov 97 printer.ml Fixed the printing of iterated polymorphic binary operators without brackets by changing dest_binop to DEST_BINARY where needed, so any instances of operators with the same name are accepted. This function DEST_BINARY is new: it works for variable infixes as well as constants. Wed 26th Nov 97 arith.ml Added: DIV_0 = |- !n. ~(n = 0) ==> (0 DIV n = 0) MOD_0 = |- !n. ~(n = 0) ==> (0 MOD n = 0) EXP_1 = |- !n. 1 EXP n = 1 DIV_LT = |- !m n. m < n ==> (m DIV n = 0) LT_EXP = |- !x m n. x EXP m < x EXP n = 2 <= x /\ m < n \/ (x = 0) /\ ~(m = 0) /\ (n = 0) Thu 20th Nov 97 basics.ml, tactics.ml Finally, I've done something about a fast save/resume. I started by creating a (term-netted) stock of theorems, optionally with names. The old COMPACT function now not only compacts theorems, but also stores them in this structure. This "theory file" is loadable and savable. Also, "prove" tries to find a theorem in the net before running the tactic to prove it; the disadvantage is that it still *creates* the tactic --- in cases where this is expensive we're going to need something like a "fun g -> tac g" wrapper. Sat 15th Nov 97 preterm.ml, sets.ml Fixed a bug pointed out by Mark Woodcock: in general set abstractions {t | P}, the only variables getting existentially quantified were those in both t *and* P, the idea being that if a variable doesn't appear in P, then it means something outside and the user doesn't want to bind it. However this is quite wrong, as pointed out by Mark's example { x | T }. So now I quantify any variables appearing in t. This meant I had to change the definition of IMAGE f s = {f x | x IN s} to an explicit existential. Really, it would only be possible to do this better using parsing in a context of meaningful (in particular bound) variables, and that creates all sorts of new problems. Probably better to accept occasional inconvenience. Tue 11th Nov 97 real.ml, int.ml Removed a few duplicated theorems. Tue 11th Nov 97 term.ml, drule.ml Chaged "paconv" to delete the env argument at toplevel. Tue 11th Nov 97 term.ml, thm.ml, bool.ml, equal.ml, basics.ml Moved syntax constructors from term.ml out to more appropriate places, since this is after all part of the core. Mon 10th Nov 97 bool.ml Put back an old optimization into PROVE_HYP, which is now completely different, so that it does no inference in the trivial case. Mon 10th Nov 97 thm.ml, equal.ml Changed the beta conversion primitive so that it only does the trivial case. This is now called BETA. The general case, BETA_CONV, is done by adding INST where needed. Mon 10th Nov 97 thm.ml, equal.ml, bool.ml Returned to old name INST despite the fact that it instantiates in assumptions, as does INST_TYPE. Profiling indicated this was safe, i.e. failure never occurred at all, let alone in a situation where failure causes some other action. Sun 9th Nov 97 term.ml, thm.ml, equal.ml bool.ml Changed the primitive basis to cut out implication as a primitive. Introduced the new inference rule DEDUCT_ANTISYM_RULE, made instantiation apply to assumptions, and moved MP, DISCH, IMP_ANTISYM_RULE and SYM from primitive to derived. The derivation of SYM is something independent from the effort of eliminating implication; I could have done it before. It doesn't appear that often in profiles, whereas TRANS does. Sun 12th Oct 97 sets.ml Strengthened FINITE_UNION to an equation, and renamed the implicational version FINITE_UNION_IMP. Wed 17th Sep 97 meson.ml Changed "Var" to "Fvar" to avoid a clash with the standard term constructors; this was also pointed out by Mark Woodcock. Wed 17th Sep 97 class.ml Deleted duplicate proof of BOOL_CASES_AX (pointed out by Mark Woodcock). Mon 8th Sep 97 wf.ml Added an "exists unique" version of the wellfounded recursion theorem. Mon 8th Sep 97 sets.ml Remarkably, I had UNIONS and INTERS defined wrong way round, and only now discovered it while trying to use them! Fixed it. Fri 5th Sep 97 wf.ml Simplified the proof that recursive existence implies wellfoundedness following a suggestion of Torkel Franzen. Wed 3rd Sep 97 gtt.ml, wf.ml, arith.ml Added a proof that the existence part of wellfounded recursion implies wellfoundedness. Needed to swap arith.ml and wf.ml in the load sequence and transfer a few theorems in order to make this cleaner. Also moved "recursion.ml" back in the load sequence so as to make "prove_recursive_functions_exist" available during wellfoundedness proof. Sat 21st Jun 97 list.ml Added FORALL_MAP = |- !P f l. FORALL P (MAP f l) = FORALL (P o f) l Fri 20th Jun 97 ind-types.ml Fixed a bug in "derive_recursion_theorem", which was failing in the case where a type had only a single constructor; the theorem was being treated as if it were a conjunction even though it wasn't. Sat 7th Jun 97 tactics.ml Modified VALID so it takes into account ASSUME-ability of conclusions of hypotheses (if you see what I mean). Since it should now be OK, inserted it into "TAC_PROOF" and "e". Sat 7th Jun 97 bool.ml Optimized PROVE_HYP so it doesn't do any work in the trivial case. This seems wise now it will get called even more often in tactic proof reconstructions. Sat 7th Jun 97 tactics.ml Modified RULE_ASSUM_TAC so that it works via a series of ASSUME_TACs; this allows the traditional HOL practice of ASSUMEing assumptions after modification. I believe now that all ways of getting a theorem onto the assumption list support this style. Fri 6th Jun 97 tactics.ml Changed b() so it doesn't back up to an empty goalstack. Tue 3rd Jun 97 tactics.ml Added checks to X_GEN_TAC and X_CHOOSE_TAC to ensure that the variables chosen aren't already free. Fri 30th May 97 tactics.ml Added "flush_goalstack". Tue 27th May 97 real.ml Added REAL_POW_LT2 = |- !n x y. ~(n = 0) /\ &0 <= x /\ x < y ==> x pow n < y pow n Fri 23rd May 97 real.ml Added REAL_POW_SUB = |- !x m n. ~(x = &0) /\ m <= n ==> (x pow (n - m) = x pow n / x pow m) and REAL_POW_NZ = |- !x n. ~(x = &0) ==> ~(x pow n = &0) Wed 21st May 97 type.ml, term.ml, preterm.ml, thm.ml Made sure that the binary type constructors are not treated as if they take pairs; after a message from the CAML list I'm not sure this is safe. Tue 20th May 97 real.ml Added REAL_OF_NUM_SUB = |- !m n. m <= n ==> (&n - &m = &(n - m)) Tue 20th May 97 meson.ml Deleted the unused internal function "bump_hol_thm" (pointed out by Michael Norrish again). Mon 19th May 97 arith.ml Added EVEN_EXP = |- !m n. EVEN (m EXP n) = EVEN m /\ ~(n = 0) and ODD_EXP = |- !m n. ODD (m EXP n) = ODD m \/ (n = 0) Sun 18th May 97 real.ml,calc_rat.ml Added SUM_HORNER = |- !f n x. Sum (0,SUC n) (\i. f i * x pow i) = f 0 + x * Sum (0,n) (\i. f (SUC i) * x pow i) and a corresponding conversion REAL_HORNER_SUM_CONV. Wed 14th May 97 meson.ml Deleted duplicate definition of POLY_ASSUME_TAC; thanks to Michael Norrish for pointing this out. Tue 13th May 97 real.ml Added REAL_POW_ONE = |- !n. &1 pow n = &1 and REAL_POW_MONO = |- !m n x. &1 <= x /\ m <= n ==> x pow m <= x pow n Mon 12th May 97 gtt.ml Moved "quot.ml" to earlier in the build sequence, in an attempt to move closer to the (unrealizable) ideal of deductive system first, theories second. Mon 12th May 97 calc_num.ml Fixed a bug in the rewrites for EXP, which was leaving "1" instead of "BIT1 _0" in the right hand side; this was then unsimplifiable by the other rewrites in certain situations. Sat 10th May 97 preterm.ml, passim Modified "overload" so it checks that the thing being overloaded (a) conforms to the type skeleton, and (b) isn't polymorphic. Also changed "overesolve" so that it will accept overloading of variables, purely so that the overloaded form (and parse status) can be used during an actual definition. Sat 10th May 97 mizar.ml Made the "given" construct use the types of variables in the thesis; even improved "consider" slightly by linking the evars and the body before typechecking. Did a few other bits of tidying up. Fri 9th May 97 parser.ml Deleted the stuff that parses types after binary operators. By construction it could never get called; types are attached to the unary subterms. Fri 9th May 97 preterm.ml, passim Added a fairly major new feature: operator overloading. The main new functions are "make_overloadable", "overload", "unoverload", "prioritize_overload" and "retypecheck", though these include various auxiliary functions. It seems to work quite well. Thu 8th May 97 tactics.ml Added "top_realgoal", which is useful for investigating goalstacks. Thu 8th May 97 gtt.ml Changed "loadt" to be less dependent on the Unix library. Now it only depends on Unix for the PID, and this could be avoided. Thu 8th May 97 mizar.ml Simplified Mizar to avoid using metavariables; hence diffuse statements are now banned. This makes lots of things much simpler. Mon 5th May 97 pair.ml, basics.ml Changed the bodies of generalized abstractions to use a new constant "GEQ" rather than equality. The problem is that with standard equality, "\(tup). T" gets its equality predicate obliterated by the default rewrite "(x = T) = x". Changed syntax functions and PAIRED_BETA_CONV correspondingly. Sun 4th May 97 arith.ml Added the definition of "measure" (natural number measures) and proved wellfoundedness. Wed 30th Apr 97 passim Changed a lot of "b => x | y"s into "if b then x else y"s. Sun 27th Apr 97 parser.ml, printer.ml Made the syntax "if b then l else r" acceptable as well as "b => l | r". Made the printer print conditionals as "if-then-else". In the long run, I might try and scrub the conditional expression; I think "if-then-else" is normally clearer. Mon 21st Apr 97 ind-defs.ml, basics.ml Moved the function "make_args" back from ind-defs into a more general place, since it's used in other places. Also tweaked it to avoid using numbers when only a single argument is required. Mon 21st Apr 97 ind-types.ml Added "prove_cases_thm" to prove the exhaustive case analysis theorems. Also changed the variable name stylization in the induction theorems to use "x" rather than "a" since the latter is already in use for the arguments. Mon 21st Apr 97 term.ml, basics.ml, bool.ml, tactics.ml Backed off the change to "variant" and instead installed an alternative "mk_primed_var" and used that in a few tactics instead of "variant". Fri 18th Apr 97 trivia.ml Deleted the constants S and K and redefined I directly. They were not useful and frequently clashed with variable names, even more troublesome given the change to variant. Fri 18th Apr 97 term.ml Changed "variant" so that it also avoids constant names. This seems useful in many situations, e.g. in problems with CHOOSE_TAC in the presence of constants with the same name as the existential variable (this was pointed out by B. Karthikeyan). Fri 18th Apr 97 class.ml Fixed a subtle bug in the tautology checker found by Claire and Tom at Glasgow. The collection of useful case splits was being accumulated at the start. However in certain composite expressions other than boolean variables, new subterms can be created by the case splits chosen by the unsubtle "deepest first" algorithm, e.g. from a term "P(x:bool)" two new terms "P T" and "P F". Fixed this by moving the "ok" accumulation into the REPEAT loop. Sun 23rd Mar 97 calc_num.ml Term-netted the general numeral arithmetic conversion to make its discrimination a bit faster. Added Karatsuba multiplication, which is a lot faster for really big numbers. It's pretty crude and could be optimized a great deal. Sat 8th Mar 97 basics.ml Slightly modified the term compactor: since "mk_mconst" recreates its types anyway, there's no point in rehashing the type. Fortunately there aren't usually all that many different polymorphic instances of constants! Also made the hash tables visible, for convenient tweaking. Sat 8th Mar 97 sets.ml Added stuff about recursive definitions over finite sets. The proofs are a slight variant of those from Ching-Tsun Chou for his HOL88 "aci" contribution. It's actually quite tricky; we copy his trick of defining a relational form including the cardinal, as in Tom Melham's original definition of CARD. Fri 7th Mar 97 sets.ml Added some trivia about finiteness being preserved by subset, union and intersection. Also added definitions of cardinal relations, but not much more. Wed 5th Mar 97 meson.ml Put an EQ-test optimization in unification and equality-under-instantiation tests. Wed 5th Mar 97 real.ml, calc_rat.ml Added a "proper" definition of summation, and a load of theorems from the old reals library. Threw away "Sum0". Wed 5th Mar 97 parser.ml Added another syntax class of "typed_appl_preterm", to allow types on subcomponents of binary operators, e.g. `x:A,y`. This has the side-effect that `x,y:A` means `x,(y:A)` not `(x,y):A`. We could change this by slightly generalizing "precedence" to take two different parsers, but the new behaviour seems at least as intuitive. Mon 3rd Mar 97 mizar.ml Corrected Mizar mode so that "endcase" does not perform an "end". This is achieved by doing a SUBGOAL_THEN followed by a "conclusion" in the global goal. There are still other things to fix here; in particular there are problems with the treatment of metavariables in several ways; best to avoid "now" for the present! Sun 2nd Mar 97 mizar.ml Changed treatment of case splits so "per cases" creates two subgoals, the first trivial, and "suppose" now *always* tries to finish the current subgoal before splitting. This seems neater than repeating each time a test for whether there were any conclusions; this was hard when using metavariables. Also fixed a bug in transitivity chaining. Sat 1st Mar 97 meson.ml Made the polymorphic duplication more generous: it was missing trivial instantiations if it could find any others. Sat 1st Mar 97 parser.ml Changed the parsing of set enumerations to parse a list of comma separated "typed_apreterms", rather than parsing a preterm then splitting pairs. The snag with that was that it was impossible to enumerate a set of pairs, since even bracketing was ignored. Sat 1st Mar 97 list.ml Added ITLIST and EX(ISTS). Sat 1st Mar 97 int.ml Corrected ARITH_TAC so that it deals correctly with non-variable terms; previously, it only worked if the ultimate constituents of the terms were atoms. Fri 28th Feb 97 mizar.ml, arith.ml, real.ml, int.ml Changed the treatment of transitivity chains to be more powerful and useful. Now they store some transitivity theorems, and these are used to connect together binary relations where the left of the second is "...". Fri 28th Feb 97 mizar.ml Generalized labelled formulas so one can label the conjuncts, separating them by "and". Also cleaned up the parser and treatment of let-bound variables, which were working out wrong when dollared. Thu 27th Feb 97 mizar.ml Fixed up Mizar mode with a parser based on the new combinators. Tue 25th Feb 97 ind-types.ml Added tools to prove injectivity and distinctness of inductive type constructors. Tue 25th Feb 97 list.ml Added FORALL Mon 24th Feb 97 recursion.ml Fixed a bug in proving recursive functions exist, so it still works even if the user only wants to justify some, not all, of the recursive functions. This is useful as a subroutine for the expansion of nested types, even if the user is unlikely to want trivially mutual definitions. Mon 24th Feb 97 lib.ml Added "set_eq", a set comparison of lists. Sun 23rd Feb 97 pair.ml Manually redid an automatic proof --- it was already highly inefficient using meson, and now with duplication of polymorphic theorems it's even worse. Sun 23rd Feb 97 tab.ml, meson.ml Deleted tableaux prover, which seems merely a diversion now meson can cope with polymorphism. Also in the interest of simplifying interfaces, threw away equality-free versions of meson and deleted the "EQ" from names. Sun 23rd Feb 97 meson.ml Fixed up meson so that polymorphic lemmas to x_MESON_TAC are automatically duplicated for relevant-looking instances in the existing assumptions. This seems crude and is probably incomplete, but it seems to make meson cope with most or all cases of polymorphic instantiation that occur in practice. Sun 23rd Feb 97 meson.ml Slightly improved the standard equality handling by adding just (x = y) /\ (x = z) ==> (y = z) instead of separate symmetry and transitivity laws. Sun 23rd Feb 97 bool.ml Fixed "is_beq", which was always returning false by comparing wrong type with `:bool`. Sun 23rd Feb 97 recursion.ml Corrected the argument reshuffling so that it works even if the recursive instances of the function are applied to fewer arguments than in the varstructs on the left. Sat 22nd Feb 97 meson.ml Installed Brand's transformation (SIAM J. Computing, vol. 4, pp. 412-430, 1975). Pending detailed comparative tests, left the old equality stuff in as an alternative. Early indications are that most cases get quicker (proofs a bit longer but found more quickly) but a few hard ones get quite a bit longer. Sat 22nd Feb 97 ind-types.ml At last, cleaned up the mutually recursive types so the postulated recursive functions have different codomain types. Reshuffled definition of sum type as we need this first. Wed 19th Feb 97 recursion.ml Reimplemented "new_recursive_definition" to work properly, even in the mutual case. Tue 18th Feb 97 parser.ml, passim Rewrote the parser in terms of nicer infix combinators. Also made the installation of user parsers more flexible (try-fail interface rather than keying on brackets). Also, treated "=" explicitly with lower precedence than conditionals. Mon 4th Nov 96 lib.ml Added multiset union and subtraction (could no doubt be done much more efficiently). Sun 3rd Nov 96 printer.ml Put in a rather ad hoc printer for program stuff. One day all this will be completely reweitten and replaced by something nice. Sun 3rd Nov 96 wf.ml Added a theorem saying that applying a measure function into a WF ordering gives a WF ordering. Sun 3rd Nov 96 tactics.ml Changed META_EXISTS_TAC to X_META_EXISTS_TAC, and provided a META_EXISTS_TAC that invents a variant of the original name. Sat 2nd Nov 96 sets.ml Beefed up "IN_ELIM_THM" with a predicate version of the GSPEC clause. Sat 2nd Nov 96 preterm.ml Corrected "undollar" so it also works inside varstructs. Fri 1st Nov 96 printer.ml Threw in printing of generalized abstractions. Really, all this stuff could be done much more neatly. Thu 31st Oct 96 ind-defs.ml Added a few extra entry points and a flag to "generalize_schematic_variables" so that the user can avoid gratuitously generalizing monotonicity hypotheses that weren't proved. This came up in defining the WP semantics of while loops; one wants to modify the monotonicity assumption to the monotonicity of the body. Wed 30th Oct 96 mizar.ml Put in a couple of pre-provers to get the common cases quickly. First try to match a conjunct of the assumptions, then try rewriting. This works out a lot quicker on average. Wed 30th Oct 96 tab.ml Added two new equality handling features to tableaux. Though incomplete, they should be pretty useful. First it uses initial equational theorems as rewrite rules. Second, it uses any clauses "v = t" or "t = v" to rewrite with "v = t" everywhere and dispose of the assumption. Finally, added a type variable renaming feature to better allow unification with type instantiation. Wed 30th Oct 96 simp.ml Added "LIMITED_REWRITE_CONV", which is useful in a few situations. Fri 25th Oct 96 bool.ml Made a few small tweaks, mainly replacing "type_of v" by "snd(dest_var v)" where "v" is a type. Thu 24th Oct 96 preterm.ml Corrected the function "typify": the environment was being used to attach types to binding instances of variables. Tue 22nd Oct 96 meson.ml, pair.ml Fixed yet another bug in the equality adder: it was creating congruence rules assuming that constants are fully applied. However after first order reduction this is not so for the I (application) combinator. Now fixed. At last, meson proves the goal in "pair.ml" that it should --- although it takes a while of course. Tue 22nd Oct 96 canon.ml, meson.ml Rewrote the FOL reduction so that it can be applied in an integrated manner to all the assumptions of a goal, rather than merely a term at a time. This fixed a few cases where meson would fail. Also made the precanonicalizer for meson simpler by using genvars for specializations. Tue 22nd Oct 96 drule.ml Fixed a bug in matching where even pure variables were treated as a higher order match, requiring all type instantiations to have been resolved. Mon 21st Oct 96 mizar.ml Finished a completely rewritten version of Mizar mode, with metavariables for diffuse statements, nested proofs, alternative prover via "by ... with ..." and various other enhancements and bugfixes. Got out the following nice proof of Tarski's fixpoint theorem: let f be A->A; assume L:antecedent; antisymmetry: (!x y. x <= y /\ y <= x ==> (x = y)) by L; transitivity: (!x y z. x <= y /\ y <= z ==> x <= z) by L; monotonicity: (!x y. x <= y ==> f x <= f y) by L; least_upper_bound: (!X. ?s:A. (!x. x IN X ==> s <= x) /\ (!s'. (!x. x IN X ==> s' <= x) ==> s' <= s)) by L; set Y_def: Y = {b | f b <= b}; Y_thm: !b. b IN Y = f b <= b by rewriting with Y_def,IN_ELIM_THM,BETA_THM; consider a such that lub: (!x. x IN Y ==> a <= x) /\ (!a'. (!x. x IN Y ==> a' <= x) ==> a' <= a) by least_upper_bound; take a; !b. b IN Y ==> f a <= b proof let b be A; assume b_in_Y: b IN Y; then L0: f b <= b by Y_thm; a <= b by b_in_Y, lub; so f a <= f b by monotonicity; hence f a <= b by L0, transitivity; end; so Part1: f(a) <= a by lub; so f(f(a)) <= f(a) by monotonicity; so f(a) IN Y by Y_thm; so a <= f(a) by lub; hence thesis by Part1, antisymmetry; end Sun 20th Oct 96 basics.ml and others Added "compact" and "COMPACT" to enforce maximal sharing on terms and types (well, not quite, because of the way mk_const works, but not bad). Made separate "prove" and "PROVE" which respectively do and don't apply the compactor. So far the results are positive but the improvement is surprisingly modest; maybe CAML does this sort of thing itself anyway. Sat 19th Oct 96 canon.ml Fixed the lambda-remover to be less zealous, otherwise there will be problems with special binder terms. Now it will just remove lambdas in beta redexes and on either side of equations. This probably covers most cases in practice. Tue 15th Oct 96 nets.ml Fixed a bug in the term nets code. This would very seldom strike; the problem was that "label_to_store" was not remembering the environment of bound variables. This means if one of them happened to be the same as a local constant, it would disable a match with an alpha-equivalent but unequal term. This showed up in matching ASSUME `sup (\x. ?n. x = f n) - e < f n` with `sup (\x. ?m. x = f m) - e < f n`. The fix is crude but adequate: replace the variable by a genvar in such cases so it couldn't possibly be the same as a local constant. Tue 15th Oct 96 arith.ml Put in the (easy) proof of DIV_MUL_LE, which has been mk_thm'd for a long time, for some reason. Tue 15th Oct 96 meson.ml Fixed the use of "FOL_CONV" at the right point in the MESON canonicalizer, i.e. after NNF conversion and CHOOSE-ing of existential variables. The practical effect of this is that we can prove Andrews's "unique fixed point" theorem in its most natural form. Mon 14th Oct 96 basics.ml, drule.ml, and passim Completely rewrote matching and higher order instantiation. Instead of the crude old 2-phase approach where the beta conversions were applied in a rather indiscriminate way, the new version has the instantiation list contain a list of pattern variables to beta-reduce. A few proofs broke under this because accidental beta conversions was relied on. The simplest solution was simply to install BETA_THM as a basic rewrite (long-term, will use BETA_CONV in the net directly, but this is OK pro tem). Don has done this anyway, i.e. replaced ABS_SIMP with BETA_THM. Rather surprisingly this change broke no proofs at all, so it's definitely a good idea. Lots of BETA_THM's deleted from rewrites all over the system. Sun 13th Oct 96 basics.ml, bool.ml Made "type_match" keep its accumulator in the visible version. This at least makes ISPECL tidier, and seems a good idea in principle. Sun 13th Oct 96 equal.ml Tweaked SUBS_CONV so that if no substitutions occur, any new hypotheses in the equational theorems get included in the result. That this was happening was an unfortunate consequence of the implementation style. It affected later tactic proofs by propagating bogus assumptions in SUBST_ALL_TAC. Sun 13th Oct 96 meson.ml Fixed a serious and longstanding bug in the MESON equality axiom producer. It was only producing equivalence axioms for equalities of a type that actually occur in the target. However in general this is not enough, e.g. "P x y /\ ~P x z ==> (y = z)" requires equality axioms for the type of "x" too! Sun 13th Oct 96 tactics.ml Put a check for boolean goal into mk_goalstate, from where it propagates out to other relevant functions like "prove". Sun 13th Oct 96 basics.ml Did a lot of cleaning up of the basic term matching function, including avoiding inefficient recalculation of the type instantiations from the term instantiations over a set of higher order pattern subterms. Also a bit of dead code removal and other minor efficiency tweaks like using "snd o dest_var" instead of "type_of" for known variables. Also an enhancement: the outer types of patterns are matched first, which in principle could lead to more matches. Another enhancement: the local constants, besides stopping instantiations of variables, are used to fix those too for the matching of patterns. This will be specially useful for Mizar, but seems a good idea anyway. Sat 12th Oct 96 basics.ml, bool.ml, drule.ml Tidied and optimized the simple stuff in "basics.ml". Moved the binder syntax back here from "bool.ml" (hence reduced the duplication in the deneralized abstraction syntax operations, which use universal quantifiers). Also moved the dreaded "mk_thm" to the start of "drule.ml". Deleted the function "beta" completely; it was only used once in "ind-types.ml". Sat 12th Oct 96 equal.ml Made a few minor tweaks, including delting BODY_CONV which was no longer used (the old canonicalizer used it). Mostly replacing "rhs" with "rand" and reorganizing type annotations more tidily. Sat 12th Oct 96 term.ml Made various minor but probably worthwhile optimizations, including the avoidance of rebuiling some paired varstructs, and slightly tidier control flow in "inst"; also the type_of tests for known variables have been optimized. Fri 11th Oct 96 tactics.ml, passim Corrected CONV_TAC so it accepts |- p rather than |- p = T even if p is an equation. Hence removed some EQT_INTROs. Fri 11th Oct 96 tactics.ml Fixed an analogous bug in DISJ_CASES_TAC; I notices how slow really big case splits were becoming. The resulting theorem was getting all prefixes of a big disjunction. Now I recall this was a fix due to TFM for a problem in the HOL88 subgoal mechanism. I guess it won't affect mine as the assumption is explicit; should think it over though. Anyway it all seems to work. Fri 11th Oct 96 canon.ml, tab.ml, meson.ml, real.ml Replaced the old complicated (but slightly more efficient) canonicalization stuff with simpler versions using rewriting. In particular, took out all the "REFUTER" stuff which is cute but too complicated. This necessitated a lot of consequential changes in files that use that stuff. Fri 11th Oct 96 tactics.ml Fixed a bizarre bug in X_CHOOSE_TAC, which was keeping the existential assumption in the new theorem. Also a less bizarre but completely stupid one in CHOOSE_THEN: the variant was being chosen avoiding the hypotheses of the existing assumption theorems, rather than the conclusion! Fri 11th Oct 96 iprover.ml, tactics.ml Moved FIRST_X_ASSUM into the main tactic file. Wed 9th Oct 96 drule.ml Beefed up the sanity check in PART_MATCH to force equality rather than alpha conversion. Wed 9th Oct 96 term.ml Separated off "pure alpha conversion" (without a preceding equality test) as "paconv". Wed 9th Oct 96 list.ml Locally changed the bound variable names of list_INDUCT inside LIST_INDUCT_TAC so that the induction subgoals have these names, as before. Tue 8th Oct 96 cprover.ml Implemented the main prover module. However it's very slow compared with the search-separated TAB_TAC. Tue 8th Oct 96 iprover.ml Fixed a bug in the unifier. Tue 8th Oct 96 tab.ml Fixed a bug resulting from the backing off of the equality stuff; a previous version of NNF_CONV had been used that took its args in a different order. Tue 8th Oct 96 list.ml Proved the remaining theorems in the modest list theory that had previously been asserted. Tue 8th Oct 96 ind-types.ml Moved the definition of lists to here, and also proved standard induction and recursion theorems for the pair type. Tue 8th Oct 96 ac.ml + passim Now we have ordered rewriting, the small performance improvement from AC_CONV et al. hardly seems worthwhile. So I removed ac.ml from the build sequence, inserting a trivial "AC" using ordered rewriting in theorems.ml, and using ordered rewriting directly in many of the proofs. Some other things, like syntax for binary operators, have been scattered in other files. The ASSOC and DISTRIB conversions have gone in canon.ml pro tem, pending a complete rewrite of that file. Mon 7th Oct 96 cprover.ml Set up basic canonicalization conversions for the classical prover and the mechanism to set it up as a prover, using a goalstate as a context. Sun 6th Oct 96 simp.ml Corrected some of the simplifier traversal strategies, and generally rewrote and optimized stuff. Sat 5th Oct 96 simp.ml Made a few minor tweaks, including optimizing and correcting the default term ordering. Sat 5th Oct 96 theorems.ml Now all the equality-free theorems here are provable by ITAUT rather than special tactic scripts, so I've done that. In the long run, we could deal with equality too. Sat 5th Oct 96 tactics.ml Fixed a bug in META_EXISTS_TAC, and a corresponding one in EXISTS_TAC; these were not instantiating the pattern argument to exists (left part of paired argument). This was causing problems with proving things involving existential quantifiers in ITAUT_xx. Fri 4th Oct 96 tactics.ml Fixed a bug in "by", which was passing the non-updated instantiation to the body of the new justification function. Fri 4th Oct 96 iprover.ml Replaced the ad hoc tautology prover by one using tactics, exploiting the metavariable feature. This also does some first order logic (without equality), modulo bugs. Fri 4th Oct 96 tactics.ml Added (as an experiment) two metavariable tactics: META_EXISTS_TAC and META_SPEC_TAC. Fri 4th Oct 96 tactics.ml Optimized UNDISCH_THEN, so now it uses the theorem directly rather than doing inference. Fri 4th Oct 96 tactics.ml Added "FIND_ASSUM" to use an assumption with given conclusion. (A more efficient alternative than using ASSUME, now to be deprecated.) Fri 4th Oct 96 bool.ml, canon.ml Corrected "is_beq" and moved it back to "bool.ml". Fri 4th Oct 96 simp.ml, tactics.ml Reversed the build order of these two files so that tactics come first. This seems better as we can then use tactics to get a simple logical theorem prover. Accordingly, moved the rewriting tactics out of "tactics.ml" and into "simp.ml". Tue 1st Oct 96 ind-types.ml Fixed a silly little bug in the "extract_arg" function; it only worked when the string of FSTs and SNDs generated by a recursive call was exactly one of FST and SND without nesting. This caused examples from Steve Brackin and Konrad Slind to fail. Now they all seem OK. It even handles Konrad's big 68000 instruction set example (418 seconds on swordfish). Tue 1st Oct 96 parser.ml Added semicolon and comma to the stock of "brackets" (single-character identifiers). Tue 1st Oct 96 parser.ml Added an option (lexquotes) to allow recognition of items in '....' as variable names, even if they don't obey the usual syntactic conventions. This was to simplify supporting Prover logs, but is arguably useful. Anyway it's switched off. Also corrected the string quotes inside quotation parser to double quotes. Mon 30th Sep 96 lib.ml Inserted a slightly seedy hack into "set_insert" and "set_merge" to make the builtin orderings return false rather than fail when given functional values. Since the built-in ordering is lexicographic on pairs, this lets us prioritize the data lists in term nets without any special ordering. Perhaps I should do this properly... Mon 30th Sep 96 nets.ml Rewrote nets completely. First, included a facility for local constants (i.e. variables not instantiable in context), and included a function to merge sets. This is more or less ripped off from Don's hol-lite code, though I don't store the local constants in the net as he does. Also made the data lists canonically sorted, so that we can encode priorities in convnets without any special measures. Mon 30th Sep 96 lib.ml Added a function "set_merge" to perform (set) union of canon-sorted sets. Mon 30th Sep 96 theorems.ml Stepped the _AC suites for conjunction and disjunction up to _ACI ones; this still gives a normalizer with ordered rewriting (I hope). Mon 30th Sep 96 simp.ml Modified the default term ordering to be a dynamic lex order. I think this is AC-compatible and efficient, but I'm not 100% sure. We'll see... Sat 28th Sep 96 wf.ml, arith.ml Fixed two broken proofs. The "arith.ml" one was the old favourite of nonconfluent rewrites. The "wf.ml" one is a bit odd; the name of a bound variable, later freed by STRIP_TAC, changed. This is presumably because I deleted the alpha conversion wrapper round REWR_CONV, since it seemed pointless. But I guess some consequences made it to the toplevel, maybe via the ind def package. Anyway, gratifyingly few breaks. Sat 28th Sep 96 simp.ml, drule.ml, iprover.ml Created "simp.ml" containing the core of a simplifier. This is crudely speaking all the "atomic" stuff, which is a generalization of basic HOL rewriting to include ordered rewriting and conditional rewriting (though without doing anything with the conditions yet except leaving them as assumptions.) The nice thing about isolating this part first is that we can integrate it with the basic HOL rewriting. This has been done, so the rewriting stuff has been taken out of "drule.ml". I've also separated off the intuitionistic tautology prover; this is a bit half-baked and deserves to be the core of a more serious prover anyway. Sat 28th Sep 96 theorems.ml, arith.ml, reals.ml, int.ml Added xxx_AC theorems for conjunction and disjunction, and for addition and multiplication of naturals, reals and integers. This is in preparation for ordered rewriting, when lots of things will become better. Fri 27th Sep 96 ind-defs.ml, lib.ml Moved "nsplit" back into lib.ml since it seems fairly general. Fri 27th Sep 96 basics.ml Made "free_in" work up to alpgha-equivalence. Wed 25th Sep 96 tactics.ml Abandoned pro tem the keeping of a series of goalstacks as completed ones were stuck there, causing problems. Wed 11th Sep 96 equal.ml Cleaned up the file a bit, e.g. removed pointless error traps for RAND_CONV and RATOR_CONV. Fri 6th Sep 96 term-nc.ml Added an optimization to "type_of" for the case "(\x. s[x]) t". Also renamed the file simply to "term.ml" since there's no foreseeable prospect of having dB terms back. Fri 6th Sep 96 type.ml Fixed a bug in "tyvarsl", which simply returned the list unchanged due to a permutation of the arguments to itlist. Fortunately this wasn't used anywhere. Now "tyvars" and "tyvarsl" are mutually recursive. Thu 15th Aug 96 tactics.ml Renamed "merge_insts" to "compose_insts" (which better suggests its sequential nature). Also corrected (?) the implementation, which was completely wrong. Thu 15th Aug 96 tactics.ml (was newtactics.ml) Fixed a moronic error in the goalstack printer which usually caused no subgoals to be printed during an interactive goalstack proof. (Basically, it used the number of current goals minus the previous number to indicate how many to print. It should be one greater than that, while being at least 1, and with no previous goal being another special case of 1.) Thu 15th Aug 96 tab.ml Replaced the new but more complicated version of tableaux (with a half-baked and incomplete equality handling system) by the old equality-free but simple version. I'll attack this again in a different way sometime. Must also try Brand/E-SETHEO transformation to MESON. Thu 15th Aug 96 Made bool_ty (":bool"), aty (":A") and bty (":B") global; removed all their local bindings. Thu 15th Aug 96 term-nc.ml,ind-defs.ml Moves "lhand" back with the other term syntax stuff. Thu 15th Aug 96 Following a suggestion of Donald Syme, deleted the file "dsyntax.ml" and inserted these syntax functions in the natural (usually) places in the development. The main defect is that some duplications are needed in the printer, which uses some derived destructors. Eventually this will be fixed when the printer prints preterms not terms (another suggestion from Don; I think Isabelle does it this way too). Thu 15th Aug 96 preterm.ml Trivial change: made "unify" curried, to use "do_list2" instead of "zip" in its implementation. Wed 14th Aug 96 parser.ml Fixed parser not to keep swallowing undefined type constructors. This introduces more context-sensitivity into parsing. However this was my original intention, since one part of the pretype parser already made this check. Mon 12th Aug 96 ind-defs.ml Made "make_args" take a letter argument, for greater flexibility. Thu 27th Jun 96 newtactics.ml Added an experimental new scheme for tactics. This retains labels, but also (1) makes the assumptions theorems, and (2) includes metavariables. There is now a move to the notion of a "refinement" on a set of subgoals (since with metavariables, the order in which goals are solved matters). This has been used to provide goalstack-like facilities in a less ad hoc fashion. The new aspects may be buggy, but it's surprisingly upward-compatible. Mon 17th Jun 96 tacticals.ml,tactics.ml I discovered SUBST_ALL_TAC was throwing away assumptions. Fixed this and related problems by adding "POP_LASSUM_TAC" and making RULE_ASSUM_TAC use it to preserve labellings. There might be a few others of these lurking. Sun 16th Jun 96 subgoal.ml Made the subgoal state a list (i.e. stack) of goalstacks. This allows one to go into a new goal state then return to the old one (but you need to "b()" explicitly; it doesn't happen just when the subgoal is proved). This is to support Mizar processing but it's handy anyway. I think hol90 does it. Sat 15th Jun 96 subgoal.ml Added a flag "valcheck" to make validity checking optional. Sat 8th Jun 96 basics.ml, printer.ml, subgoal.ml, tacticals.ml, tactics.ml, meson.ml Introduced labelled assumptions in tactics. The main reason was to achieve a new level of integration between tactics and MESON proofs. However labels seem useful for other reasons. The main additions are "LABEL_TAC" to perform a labelled "ASSUME_TAC", and "USE_ASSUM" to grab the assumption with a given label. The goal prettyprinter puts any labels in parentheses at the right of the term, which seems quite close to mathematics books, and doesn't disturb the formatting. While tinkering with tactics, I also grabbed the chance to add "UNDISCH_THEN" which I've always wanted, and added a generic function "ASM_TAC" to augment the theorem list given to a tactic with the assumptions of the goal. In fact, if we called it just "ASM" then "ASM REWRITE_TAC[]" and "ASM_REWRITE_TAC[]" would be the same and we could avoid all those extra functions. Perhaps the HOL world isn't quite ready for that, though. Since I can't use my pet "W(ACCEPT_TAC o mk_thm)" any more, I decided to add "CHEAT_TAC"! Anyway, no proofs broke, and there are some big ones now. Of course, there was no reason why they should. Fri 7th Jun 96 ind-defs.ml Corrected a non-PE'd quotation (just `T` fortunately) which somehow got left in. Sat 25th May 96 calc_rat.ml Added "REAL_SUM0_CONV" to evaluate summations. I don't really know why I put it in this file. Fri 24th May 96 calc_rat.ml Included "REAL_RAT_DIV_CONV" in the overall conversion; it had been inadvertently left out. Also optimized REAL_RAT_DIV_CONV to fail quickly when given a canonical rational. Wed 22nd May 96 tactics.ml Added "ABBREV_TAC", which I found quite useful in the reals library in HOL88. I've now come across a use for it, and in general, I think it's a worthy inclusion. Wed 22nd May 96 calc_num.ml Added a "NUMERAL" on the right-hand side of the first ARITH_SUB theorem! Tue 21st May 96 tab.ml Put proper initialization of "cstore" in "fol_of_const" and friends, so that a call to "reset_cstore" isn't necessary at the start. This bug was pointed out by Donald Syme. Mon 20th May 96 dsyntax.ml Added "is_beq" test for Boolean equations (iffs). Sun 19th May 96 equal.ml Fixed up "BINOP_CONV" to do what it should do, i.e. ignore the operator. (The old one checked it against an additional argument). Added "DEPTH_PRED_CONV" to apply a conversion to atoms in pred. calc. Fri 26th Apr 96 calc_num.ml Added "num_CONV"; not before time! Now it's derived, of course. Sat 20th Apr 96 printer.ml Added printing of a space after the ";" and "," separators for lists and sets. Sat 20th Apr 96 dsyntax.ml, equal.ml Added "is_cons" (forgotten earlier) and a conversion "LIST_CONV" to apply a conversion to each element of a list. Sun 31st Mar 96 printer.ml Put special syntax for list printing in; also fixed a bug in printing of term sequences (it always used "," regardless of the separator argument, and due to a syntax ambiguity in CAML, it wasn't dealing with empty lists correctly). Sun 31st Mar 96 dsyntax.ml Put in syntax functions for lists: mk_cons, dest_cons, mk_list, mk_flist, dest_list and is_list. Sun 31st Mar 96 real.ml Took the rewrite-based natural number arithmetic conversions out; now the new, more efficient ones are used. This file now does seem to go through a bit faster (2:25 rather that 2:50 real time for the REAL_ARITH_TAC applications). Sun 31st Mar 96 calc_num.ml Finished a complete suite of numeric calculation routines. They're OK, but really multiplication should be O(n^log_2(3)) optimized. Anyway, the performance is not too disastrous, e.g. `2 EXP 1000` takes under 7 seconds; `(2 EXP (4 + 5) * 2) DIV (3 EXP (8 MOD 4))` takes 0.35 seconds (on woodcock). Note the "n * 1 = n" traps in multiplication; given these, then denormalization will never occur, per construction (or so I hope). Sun 31st Mar 96 arith.ml Added theorems about the uniqueness of DIV and MOD, and derived some earlier theorems from them. Sat 30th Mar 96 arith.ml calc_num.ml Separated off the arithmetic rewrites into a separate file "calc_num.ml". Began augmenting these crude rewrites with some more efficient conversions. Sat 30th Mar 96 dsyntax.ml arith.ml parser.ml printer.ml Wrote versions of numeral conversions using Valerie's bignums. Just in case, added "small" versions for type "int"; maybe these should be scrubbed in fact. Moved the numeral constructors forward into "arith.ml" to get the benefit of a bit of PE. (The destructors are needed in the printer at present.) Fixed up the parser and printer to use bignums; changing pmk_numeral. Sat 30th Mar 96 - Moved over to version 0.71 of CAML without obvious problems. Took the chance to set up a preload of both unix and bignum stuff in "my_little_caml". Sat 30th Mar 96 real.ml Fixed a thinko in "REAL_ATOM_NORM_CONV". I'd defined an optimized conversion internally to avoid pointless "&0 = a" -> "&0 = a - &0" canonicalization, but then never used it. This ought to make the real arithmetic stuff marginally faster on average. Fri 29th Mar 96 arith.ml, parser.ml, perterm.ml, dsyntax.ml, num.ml Put an extra tag "NUMERAL" round all bitwise numerals. This is to fix an old bug where one numeral is actually a subterm of another, leading to bizarre effect such as rewrites with |- 1 = SUC 0 to "2". The change was pretty trivial to make, fixing all the numeral manipulating syntax functions, parser etc. and throwing in NUMERAL as an additional rewrite to derive some theorems. All the arithmetic rewrites now have the extra "numeral" tag. NB! We consistently use the numeral tag for zero too (which therefore ceases to be a constant). This may require a modest hack when we get round to defining "new_recursive_definition" properly. As it is, we now have "_0" for the constant (which is needed as a foundation for the numerals). Thu 28th Mar 96 thm.ml Hey -- our first unsoundness bug since teething troubles over prelogic! This was pointed out by Tom: in the definition rules there's a requirement not to have type variables free in the term being defined which aren't free in the new constant. For example: let th0 = new_definition `X = !x:A. !y. x = y`;; let th1 = INST_TYPE [`:one`,`:A`] th0 and th1' = INST_TYPE [`:bool`,`:A`] th0;; let th2 = TRANS (SYM th1) th1';; let th3 = ONCE_REWRITE_RULE[one] th2;; let th4 = REWRITE_RULE[] th3;; let th5 = SPECL [`T`;`F`] th4;; let th6 = REWRITE_RULE[] th5;; Fixed this by a check in "new_basic_definition" that all type variables in the predicate are in the type of the constant to be defined. I think this is the right thing: it means that type instantiation of the predicate is reflected in a change to the constant. But should think this over carefully one day... Thu 28th Mar 96 lib.ml Put a "subset" function in. About time; see above for the reason. It would be more efficient to sort first, but I don't really mind... Thu 21st Mar 96 meson.ml Fixed a very strange bug in the production of congruence rules for EQ_MESON. It was completely wayward, producing something not in clausal form. Obviously I'd just assumed this worked without even the most rudimentary testing. Wed 20th Mar 96 meson.ml Fixed a gross blunder in "fol_of_hol_clauses". It was taking as the set of local constants the hypothesis frees of the head theorem, but now takes the hypothesis frees of the theorem it's actually doing. What was I thinking? Sat 17th Feb 96 drule.ml Removed "part_match_error". I've no idea how that got there. Anyway the compiler couldn't infer the types ("I" in a ref) so it spoiled my profile run. Wed 14th Feb 96 gtt.ml Moved "trivia" back before "canon", since we might as well use I_THM for first order reduction. Wed 14th Feb 96 canon.ml, meson.ml, tab.ml Wrote a conversion FOL_CONV to produce a better FOL reduction (Donald Syme was already finding examples where the naive apporach failed). Basically it finds the minimal application level for each "constant", then fills out any greater levels with explicit invocations of the "apply" operator (we use I for this -- probably should use a separate constant just in case the original formula involves "I"...) This is now used by tableaux and MESON. Wed 14th Feb 96 canon.ml Put in an extra conversion for REFUTE and CNF_REFUTE which is applied after the reduction to NNF, but before any splitting etc. The idea is to allow a superior massage into first order form. Wed 14th Feb 96 printer.ml Put a break after each assumption in "print_thm". Tue 13th Feb 96 real.ml Proved all the trivial lemmas leading up to the arithmetic decision procedure, which had previously been mk_thm'ed (they all turned out to be true!) Really, should go over these again when the equality handling in TAB_TAC improves; it should do almost all of them automatically. Tue 13th Feb 96 canon.ml, tab.ml, meson.ml, real.ml, int.ml Added conversions EQ_ABS_CONV and UNLAMB_CONV. The latter tries to eliminate lambda-terms from formulas as input to first order automation tools. Now since the elimination of a lambda requires the resulting term to be further processed (to get it into NNF and remove any further lambdas), the NNF_CONV function has been modified to do a retraversal after calling "baseconv". Now to stop this, the conversion given must be NO_CONV, not REFL. This has been appropriately fixed up. The extra baseconv argument has been retained in CNF_REFUTE too. Now tableaux and MESON use UNLAMB_CONV in their NNF steps. Mon 12th Feb 96 thm.ml Put an extra test in BETA_CONV to avoid calling "vsubst" in the special situation where the abstraction variable and argument are the same, i.e. "(\x. t[x]) x". This was intended to make let-elimination more efficient in certain situations, but should have other efficiency payoffs since these trivial beta-redexes often arise in higher order matching, and in SPEC_ALL. Sun 11th Feb 96 parser.ml, printer.ml, dsyntax.ml, pair.ml Changed the representation of lets so that "let x = s and y = t in u[x,y]" becomes "LET (\x y. LET_END (u[x,y])) s t"; the new constant "LET_END" is used to flag the end of the let-terms in case "u" is itself an abstraction. Put in a "dest_let" and "is_let"; the former is now used in the printer. Sun 11th Feb 96 printer.ml Increased box limit (after which elision starts) from default of 35 to 100. Sun 11th Feb 96 parser.ml, printer.ml Added parsing support (actually just a correct "pmk_let" constructor; the parser was already OK) and printing support for let-expressions. Fri 9th Feb 96 equal.ml Fixed a fumble in TOP_SWEEP_QCONV, which had TOP_DEPTH_QCONV in place of a recursive call. Thanks to Don Syme for pointing this out. Thu 8th Feb 96 realax.ml Filled in the remaining gaps in the real construction: construction of multiplicative inverse and completeness transfer from ":hreal". Fri 2nd Feb 96 arith.ml Added a proper derivation of the existence of DIV and MOD, and defined them by constant specification. The proof was basically the same as (Tom's?) HOL88 one. Tue 12th Dec 95 meson.ml, arith.ml, realax.ml Added versions of MESON_TAC which take theorems and allow throwing in of the assumptions, by analogy with REWRITE_TAC and ASM_REWRITE_TAC. These are much more convenient to use, and instances in other files have been updated. Tue 12th Dec 95 canon.ml Added "CONV_OF_PROVER", useful for more palatable forms of MESON_TAC, TAB_TAC and no doubt others in the future. Fri 8th Dec 95 fol.ml Fixed a bug in the addition of equality axioms; it was leaving implications in the supposed clausal form of the congruence rules. Wed 6th Dec 95 canon.ml Used a different version of EXISTS_UNIQUE_THM to eliminate unique-existence; this should make the resulting proofs easier. Wed 6th Dec 95 fol.ml Fixed a bug in the HOL translation of MESON proofs. The local instantiations saved in the proof tree were merely being unioned in, but it can happen that the toplevel instantiations change them (because some variable free in the, err, residue, gets instantiated "right" in the proof. Wed 6th Dec 95 canon.ml Added a preprocessing phase to NNF_CONV which eliminates any logical constants T and F used in combination; it also makes a few other handy simplifications and expands unique-existence instances. The REFUTE function checks whether it's already got "F", just in case that would stump the later function! Tue 5th Dec 95 basics.ml, drule.ml, ac.ml, ind-types.ml Introduced an extra argument "local constants" for term_match. This is instantiated to the variables free in both hyp and concl of the theorem when it's used in PART_MATCH. The main idea is to inhibit impossible higher order matches, and so improve efficiency. However it may lead to usefully more rapid failures in other no-match situations too. Sun 3rd Dec 95 class.ml Took out duplicate SELECT_AX. How did that get there? Tue 28th Nov 95 printer.ml Fixed up the printing of negation properly: it was giving it precedence over infixes --- actually HOL88 does this for high-binding infixes, meaning basically not the logical operations. Perhaps I should introduce precedence for prefixes and follow suit. Also generalized printer to arbitrary prefixes (though in fact we don't have any besides negation yet!) Sun 26th Nov 95 fol.ml Added a new version of MESON_TAC which throws in the equality axioms to the proof search. The initial results are not encouraging... For example Pelletier number 49 takes 6 minutes! But some of the others aren't too bad... Sun 26th Nov 95 sets.ml More set theory stuff ported, using automated proofs; however we still specify by trial and error which equality theorems etc. to throw in. There's also a slight problem over when we want the extra theorems to get rewritten with the definitions (use MP_TAC first when we do, rather than giving it in the theorem list argument). Sun 26th Nov 95 It was just getting too tedious being unable to interrupt things, now we have all these first order automation tools, so... Made a signal call so that SIGINTs raise a new exception "Interrupt". Then trawled through systematically taking out all "with _ ->" traps. Most can be replaced by "with Failure _ ->", since practically every function used is one of ours (except arithmetic, and I think that just wraps without failing). A few need to be treated more carefully (Unchanged exceptions and Cut too in the Prolog engine). So, after a small number of fixes, everything seems to work again, and now interrupts will always(?) propagate. Really, it would be preferable to have a separate signal mechanism, but... Fri 24th Nov 95 fol.ml Modified the tableau procedure to use a similar "offset" technique to MESON in order to create temporary variables. This is much better than inventing stacks of genvars of its own, each with a HOL analog... Thu 23rd Nov 95 fol.ml Added a depth bound (as opposed to inference bound) option to MESON_TAC. Of course it's important that this switches off the divide-and-conquer optimization; this is forced even if the user forgets. We also disable the size consideration in the continuation cacher, to avoid uneccessary conservatism. Wed 22nd Nov 95 fol.ml Added MESON_TAC. It seems to work reasonably well. Tue 21st Nov 95 canon.ml Added conversions for DNF and CNF. Also added the wrapper CNF_THEN_REFUTE, which is a bit more efficient than just applying the conversion at toplevel. Tue 21st Nov 95 ac.ml Added DISTRIB_CONV and BODY_CONV. Tue 21st Nov 95 equal.ml, ac.ml, drule.ml Introduced COMB2_QCONV and COMB_QCONV; renamed COMB_CONV2 to COMB2_CONV everywhere. Mon 20th Nov 95 equal.ml, real.ml Introduced BINOP_CONV as an analog to SUB_CONV which is what I thought it was; the old one is renamed to DEPTH_BINOP_CONV and the two instances in real.ml changed. Mon 20th Nov 95 gtt.ml Fixed the filter to use suffix of pathname only when creating its temporary filename, otherwise it tries to deal with a nonexsistent directory. Mon 20th Nov 95 parser.ml Modified set enumeration parsing to use "INSERT" and "EMPTY", rather than write down an abstraction term explicitly. Mon 20th Nov 95 int.ml Modified the integer decision procedure to do discretization properly, i.e. to force NNF and hence fix the signs of all the inequalities first. The overhead shouldn't be too large since the NNF run introduced by the reals decision procedure will be trivial in most cases (the refuter will give a toplevel double negation, and the remainder is already in NNF). Mon 20th Nov 95 basics.ml Modified "find_terms" to use "insert" rather than "::", and hence produce a set. Mon 20th Nov 95 real.ml Added a quick prepass to REAL_ARITH_CONV to eliminate trivial quantifiers; this can get a few more things into prenex universal form, though it's probably quite useless in practice. Renamed "REAL_ARITH_CONV" to "REAL_ARITH" and put in a separate EQT_INTRO version called "REAL_ARITH_CONV" (just for consistency; it wasn't a conversion). Sun 19th Nov 95 real.ml Altered REAL_ARITH_CONV to force prenex form, and hence avoid failing where a trivial transformation would bring the theorem into universal form, e.g. "!x. x < y \/ !z. z <= y ==> z <= x". Sun 19th Nov 95 canon.ml Added a conversion for prenexing (assuming there are no equivalences or conditionals). Sun 19th Nov 95 int.ml Put in a decision procedure for the naturals; a wrapper for the one for integers. Sun 19th Nov 95 filter.c Tweaked the filter again; it still wasn't giving the right line numbers! Sun 19th Nov 95 real.ml Made sure all quotations get evaluated at load time. This is not just an efficiency tweak -- if a different interface map is in action at runtime, we get a change in behaviour! This was exhibited in the integer decision procedure. Sun 19th Nov 95 int.ml Added a simple decision procedure for linear arithmetic over the integers. This just hacks the term down to the corresponding real fact (we assume it's universal in these procedures anyway) then calls REAL_ARITH_CONV. Sun 19th Nov 95 gtt.ml Made the filter pick a temporary filename which depends on the PID and the original filename -- this should avoid the problem of contention between parallel GTT sessions on the same machine, and give more helpful error messages. Sun 19th Nov 95 thm.ml Made benign type redefinitions acceptable, by the crude device of inserting a cache of previous calls and results. Sat 18th Nov 95 ind-defs.ml INDUCT_THEN and INDUCT_TAC added. These are similar, but not the same, as the ones in the HOL88/hol90 library. In particular, there is no distinction made between "hypotheses" and "side-conditions", and so no call for two separate theorem continuations. Sat 18th Nov 95 sets.ml Wrote simple tactic to reduce the very elementary reasoning to FOL, then call TAB_TAC; hence automatically proved a reasonable number of the theorems in the HOL88 pred_sets library. Sat 18th Nov 95 thm.ml, bool.ml, drule.ml Made benign term definitions acceptable (up to alpha conversion; this is necessary since pairs and set abstractions may introduce different bound variable names on different occasions); also introduced storage of definitions (which was not done before, despite the presence of the variable "the_definitions"!) Fri 17th Nov 95 ind-defs.ml, num.ml, ind-types.ml Changed "new_inductive_definition" to split up the three conjuncts into separate theorems. Modified the two existing uses in the core. Fri 17th Nov 95 sets.ml [new file] Basic definitions for set theory added. Fri 17th Nov 95 parser.ml, preterm.ml, printer.ml Included parsing and printing support for enumerated and generalized set specifications; also added to "preterm.ml" a few preterm syntax functions that were needed and may be useful elsewhere. Fri 17th Nov 95 arith.ml, bool.ml, int.ml, pair.ml, realax.ml, trivia.ml, wf.ml Rationalized the infix precedences as follows (arithmetic operators include their analogs in other number systems and in set theory). They're right-associative, unless marked "L". These seem the most sensible; e.g. x - y + z and x - y - z work as expected. I'm not quite sure where pairing (,) should go really... 2 = 4 ==> 6 \/ 8 /\ 10 general_equivalence_relation 11 IN 12 < <= >= > general_order_relation SUBSET PSUBSET 14 , 16 + UNION 18 L - DIFF 20 * INTER 21 INSERT DELETE 22 L / DIV MOD 24 L pow EXP 26 o Fri 17th Nov 95 lib.ml Put in a version of "map" with a more intuitive evaluation order (left to right). Fri 17th Nov 95 fol.ml Put continuation cacheing into tableau prover (so repeated attempts will immediately fail). On big examples this seems quite useful; e.g. Andrews's challenge now takes 89.77 seconds instead of 130.23. Pelletier 43 now takes 2.21 seconds instead of 13.50. The other Pelletier examples are all pretty much the same. Fri 17th Nov 95 filter.c Changed the filter to echo all newlines (and to flush stdout after each newline, rather than after each ";;"). This makes the error messages from CAML better in two respects: (1) the line numbers are right; (2) you don't get lots of previous phrases spat out at you. Fri 17th Nov 95 real.ml Fixed a bug in the cacher's translator, which did not distribute negations over the final term in a sum. Thu 16th Nov 95 lib.ml Added a timing function. Thu 16th Nov 95 lib.ml Added "report" function and verbose/terse flag; modified "warn" to use it. Thu 16th Nov 95 fol.ml Integrated a simple tableau prover based on Lean-TAP into HOL. The preliminary results are quite good; e.g. it solves most of the equality-free Pelletier problems fairly quickly. Andrews's challenge takes about 2 minutes. Wed 15th Nov 95 real.ml Cleaned up the decision procedure REAL_ARITH_CONV and modified it to use the refuter in "canon.ml" rather than its own bespoke code. Also added intelligent cacheing of atom normaliztion theorems (i.e. it remembers negations of previous atoms --- this seemed easier than canonicalizing prior to the NNF conversion, and probably as efficient, on average, since it might on occasion catch complicated duplicates). Wed 15th Nov 95 quot.ml The "lift_theorem" function in the quotient stuff has been rewritten: given the new higher order rewriting, it's trivial. The left and right sides of the derived theorems returned by "lift_function" have been swapped; this seems more sensible. Wed 15th Nov 95 drule.ml, arith.ml, realax.ml Changed PART_MATCH to attempt a crude (but adequate) preservation of bound variable names. Really, all the higher order matching, BETA_VAR and this should be more neatly integrated in PART_MATCH, which is the practically universal entry point. There are further things one would like, e.g. the recognition that "f" and "f'" should be changed in parallel -- but where do we stop? A few consequentially broken proofs fixed; broken mainly because INDUCT_TAC now (correctly!) preserves the bound variable name in the goal rather than replacing it by "n". Mon 13th Nov 95 basics.ml Substantially enhanced higher order matching so that (1) Manifestly first order bits are handled first, which may increase the stock of determinate variables available later; this is sometimes useful; e.g. RULE_INDUCT_TAC has a more natural coding; and (2) More general patterns are allowed, of the form P (t1[x1,..,xn]) ... (tm[x1,..,xn]), where the x1,..,xn are all determinate. This is useful for doing some transfer stuff (!n. P(&n) = ...) etc. and occasionally in other places too. Mon 13th Nov 95 term-nc.ml Put trivial PE into "inst"; it returns the identity function if the instantiation list is empty. Thu 9th Nov 95 filter.c Fixed filter to ignore interrupts, and also to insert "let it = " before any toplevel phrases which are not declarations. Thu 9th Nov 95 nets.ml, arith.ml, realax.ml Reversed the appending in "follow" to favour specificity in matches; that is "SUC n" will (if matchable) be put ahead of "m" in the list of possible matches. This seems a more sensible defalt. A few consequentially broken proofs fixed. Wed 1st Nov 95 theorems.ml Cleaned up a few proofs, now that higher order matching bugs are cleared. Tue 31st Oct 95 printer.ml Terms and types are now printed with backquotes, not double quotes. Tue 31st Oct 95 parser.ml "lextoken" now uses consing followed by final reverse; probably a bit better than iterative snoc-ing. Tue 31st Oct 95 drule.ml, ind-defs.ml Moved the derived congruence rules back into drule.ml. Tue 31st Oct 95 parser.ml Put in faster character disrimination based on lookup table. (Be nicer to have bitwise operations for this as some aren't mutually exclusive). Mon 30th Oct 95 lib.ml Slightly less lamentable implementation of "explode" substituted. (The old one kept taking the `tail' of the string, which involved creating new copies!) Mon 30th Oct 95 lib.ml,preterm.ml,subgoal.ml The "warn" function added, and the two warnings to date modified to use it. Mon 30th Oct 95 preterm.ml "type_of_pretype" and "term_of_preterm" tweaked so they will invent names for STVs, but issue a warning. Sun 29th Oct 95 subgoal.ml Inserted a warning into "expand" if the tactic does not change the goal state.