# HG changeset patch # User skalberg # Date 1109850181 -3600 # Node ID 8d8c70b41babae682688db8d32aa12567195a025 # Parent 1b3115d1a8df6053092aebc417436c0f1ef36088 Move towards standard functions. diff -r 1b3115d1a8df -r 8d8c70b41bab NEWS --- a/NEWS Thu Mar 03 09:22:35 2005 +0100 +++ b/NEWS Thu Mar 03 12:43:01 2005 +0100 @@ -10,6 +10,37 @@ Library.OPTION: Isabelle now uses the standard option type. The functions the, is_some, is_none, etc. are still in Library, but the constructors are now SOME and NONE instead of Some and None. + They throw the exception Option. + +* The exception LIST is no more, the standard exceptions Empty and + Subscript, as well as Library.UnequalLengths are used instead. This + means that function like Library.hd and Library.tl are gone, as the + standard hd and tl functions suffice. + + A number of functions, specifically those in the LIBRARY_CLOSED + signature, are now no longer exported to the top ML level, as they + are variants of standard functions. The following suggests how + one can translate existing code: + + the x = valOf x + if_none x y = getOpt(x,y) + is_some x = isSome x + apsome f x = Option.map f x + rev_append xs ys = List.revAppend(xs,ys) + nth_elem(i,xs) = List.nth(xs,i) + last_elem xs = List.last xs + flat xss = List.concat xss + seq fs = app fs + partition P xs = List.partition P xs + filter P xs = List.filter P xs + mapfilter f xs = List.mapPartial f xs + + The final four functions, take, drop, foldl, and foldr, are somewhat + more complicated (especially the semantics of take and drop differ + from the standard). + + A simple solution to broken code is to include "open Library" at the + beginning of a structure. Everything after that will be as before. * Document preparation: new antiquotations @{lhs thm} and @{rhs thm} printing the lhs/rhs of definitions, equations, inequations etc. diff -r 1b3115d1a8df -r 8d8c70b41bab TFL/casesplit.ML --- a/TFL/casesplit.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/TFL/casesplit.ML Thu Mar 03 12:43:01 2005 +0100 @@ -207,7 +207,7 @@ let val (n,ty) = Term.dest_Free x in (if vstr = n orelse vstr = Syntax.dest_skolem n then SOME (n,ty) else NONE ) - handle LIST _ => NONE + handle Fail _ => NONE (* dest_skolem *) end; val (n,ty) = case Library.get_first getter freets of SOME (n, ty) => (n, ty) @@ -311,7 +311,7 @@ raise ERROR_MESSAGE "splitto: cannot find variable to split on") | SOME v => let - val gt = Data.dest_Trueprop (nth_elem(0, Thm.prems_of th)); + val gt = Data.dest_Trueprop (List.nth(Thm.prems_of th, 0)); val split_thm = mk_casesplit_goal_thm sgn v gt; val (subthms, expf) = IsaND.fixed_subgoal_thms split_thm; in @@ -340,7 +340,7 @@ the well-foundness conditions have been solved. *) local fun get_related_thms i = - mapfilter ((fn (r,x) => if x = i then SOME r else NONE)); + List.mapPartial ((fn (r,x) => if x = i then SOME r else NONE)); fun solve_eq (th, [], i) = raise ERROR_MESSAGE "derive_init_eqs: missing rules" diff -r 1b3115d1a8df -r 8d8c70b41bab TFL/post.ML --- a/TFL/post.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/TFL/post.ML Thu Mar 03 12:43:01 2005 +0100 @@ -46,7 +46,7 @@ *--------------------------------------------------------------------------*) fun termination_goals rules = map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop) - (foldr (fn (th,A) => union_term (prems_of th, A)) (rules, [])); + (Library.foldr (fn (th,A) => union_term (prems_of th, A)) (rules, [])); (*--------------------------------------------------------------------------- * Finds the termination conditions in (highly massaged) definition and @@ -105,7 +105,7 @@ | _ => r RS P_imp_P_eq_True (*Is this the best way to invoke the simplifier??*) -fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L)) +fun rewrite L = rewrite_rule (map mk_meta_eq (List.filter(not o id_thm) L)) fun join_assums th = let val {sign,...} = rep_thm th @@ -207,7 +207,7 @@ -- Lucas Dixon, Aug 2004 *) local fun get_related_thms i = - mapfilter ((fn (r,x) => if x = i then SOME r else NONE)); + List.mapPartial ((fn (r,x) => if x = i then SOME r else NONE)); fun solve_eq (th, [], i) = raise ERROR_MESSAGE "derive_init_eqs: missing rules" @@ -226,9 +226,9 @@ val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o HOLogic.mk_Trueprop) eqs fun countlist l = - (rev o snd o (foldl (fn ((i,L), e) => (i + 1,(e,i) :: L)))) ((0,[]), l) + (rev o snd o (Library.foldl (fn ((i,L), e) => (i + 1,(e,i) :: L)))) ((0,[]), l) in - flat (map (fn (e,i) => solve_eq (e, (get_related_thms i rules), i)) + List.concat (map (fn (e,i) => solve_eq (e, (get_related_thms i rules), i)) (countlist eqths)) end; end; diff -r 1b3115d1a8df -r 8d8c70b41bab TFL/rules.ML --- a/TFL/rules.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/TFL/rules.ML Thu Mar 03 12:43:01 2005 +0100 @@ -133,7 +133,7 @@ fun FILTER_DISCH_ALL P thm = let fun check tm = P (#t (Thm.rep_cterm tm)) - in foldr (fn (tm,th) => if check tm then DISCH tm th else th) + in Library.foldr (fn (tm,th) => if check tm then DISCH tm th else th) (chyps thm, thm) end; @@ -658,7 +658,7 @@ in (ants,get_lhs eq) end; -fun restricted t = is_some (S.find_term +fun restricted t = isSome (S.find_term (fn (Const("Wellfounded_Recursion.cut",_)) =>true | _ => false) t) @@ -785,7 +785,7 @@ val dummy = print_cterms "TC:" [cterm_of sign (HOLogic.mk_Trueprop TC)] val dummy = tc_list := (TC :: !tc_list) - val nestedp = is_some (S.find_term is_func TC) + val nestedp = isSome (S.find_term is_func TC) val dummy = if nestedp then say "nested" else say "not_nested" val dummy = term_ref := ([func,TC]@(!term_ref)) val th' = if nestedp then raise RULES_ERR "solver" "nested function" @@ -808,7 +808,7 @@ (prover names) (empty_ss addsimps [cut_lemma'] addeqcongs congs) ctm val th2 = equal_elim th1 th in - (th2, filter (not o restricted) (!tc_list)) + (th2, List.filter (not o restricted) (!tc_list)) end; diff -r 1b3115d1a8df -r 8d8c70b41bab TFL/tfl.ML --- a/TFL/tfl.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/TFL/tfl.ML Thu Mar 03 12:43:01 2005 +0100 @@ -201,8 +201,8 @@ fun mk_pat (c,l) = let val L = length (binder_types (type_of c)) fun build (prfx,tag,plist) = - let val args = take (L,plist) - and plist' = drop(L,plist) + let val args = Library.take (L,plist) + and plist' = Library.drop(L,plist) in (prfx,tag,list_comb(c,args)::plist') end in map build l end; @@ -336,7 +336,7 @@ val dummy = map (no_repeat_vars thy) pats val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats, map (fn (t,i) => (t,(i,true))) (enumerate R)) - val names = foldr add_term_names (R,[]) + val names = Library.foldr add_term_names (R,[]) val atype = type_of(hd pats) and aname = variant names "a" val a = Free(aname,atype) @@ -433,13 +433,13 @@ end; -fun givens pats = map pat_of (filter given pats); +fun givens pats = map pat_of (List.filter given pats); fun post_definition meta_tflCongs (theory, (def, pats)) = let val tych = Thry.typecheck theory val f = #lhs(S.dest_eq(concl def)) val corollary = R.MATCH_MP Thms.WFREC_COROLLARY def - val pats' = filter given pats + val pats' = List.filter given pats val given_pats = map pat_of pats' val rows = map row_of_pat pats' val WFR = #ant(S.dest_imp(concl corollary)) @@ -499,7 +499,7 @@ val tych = Thry.typecheck thy val WFREC_THM0 = R.ISPEC (tych functional) Thms.WFREC_COROLLARY val Const("All",_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0 - val R = Free (variant (foldr add_term_names (eqns,[])) Rname, + val R = Free (variant (Library.foldr add_term_names (eqns,[])) Rname, Rtype) val WFREC_THM = R.ISPECL [tych R, tych g] WFREC_THM0 val ([proto_def, WFR],_) = S.strip_imp(concl WFREC_THM) @@ -540,7 +540,7 @@ prths extractants; ()) else () - val TCs = foldr (gen_union (op aconv)) (TCl, []) + val TCs = Library.foldr (gen_union (op aconv)) (TCl, []) val full_rqt = WFR::TCs val R' = S.mk_select{Bvar=R1, Body=S.list_mk_conj full_rqt} val R'abs = S.rand R' @@ -622,7 +622,7 @@ let val (qvars,body) = S.strip_exists tm val vlist = #2(S.strip_comb (S.rhs body)) val plist = ListPair.zip (vlist, xlist) - val args = map (fn qv => the (gen_assoc (op aconv) (plist, qv))) qvars + val args = map (fn qv => valOf (gen_assoc (op aconv) (plist, qv))) qvars handle Option => sys_error "TFL fault [alpha_ex_unroll]: no correspondence" fun build ex [] = [] @@ -682,7 +682,7 @@ rows = map (expnd c) rows}) (U.zip3 new_formals groups constraints) val recursive_thms = map mk news - val build_exists = foldr + val build_exists = Library.foldr (fn((x,t), th) => R.CHOOSE (tych x, R.ASSUME (tych t)) th) val thms' = ListPair.map build_exists (vexl, recursive_thms) @@ -698,7 +698,7 @@ let val tych = Thry.typecheck thy val ty_info = Thry.induct_info thy in fn pats => - let val names = foldr add_term_names (pats,[]) + let val names = Library.foldr add_term_names (pats,[]) val T = type_of (hd pats) val aname = Term.variant names "a" val vname = Term.variant (aname::names) "v" @@ -733,7 +733,7 @@ in fun build_ih f P (pat,TCs) = let val globals = S.free_vars_lr pat - fun nested tm = is_some (S.find_term (curry (op aconv) f) tm) + fun nested tm = isSome (S.find_term (curry (op aconv) f) tm) fun dest_TC tm = let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm)) val (R,y,_) = S.dest_relation R_y_pat @@ -762,7 +762,7 @@ fun build_ih f (P,SV) (pat,TCs) = let val pat_vars = S.free_vars_lr pat val globals = pat_vars@SV - fun nested tm = is_some (S.find_term (curry (op aconv) f) tm) + fun nested tm = isSome (S.find_term (curry (op aconv) f) tm) fun dest_TC tm = let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm)) val (R,y,_) = S.dest_relation R_y_pat @@ -795,7 +795,7 @@ let val tych = Thry.typecheck thy val antc = tych(#ant(S.dest_imp tm)) val thm' = R.SPEC_ALL thm - fun nested tm = is_some (S.find_term (curry (op aconv) f) tm) + fun nested tm = isSome (S.find_term (curry (op aconv) f) tm) fun get_cntxt TC = tych(#ant(S.dest_imp(#2(S.strip_forall(concl TC))))) fun mk_ih ((TC,locals),th2,nested) = R.GENL (map tych locals) @@ -832,7 +832,7 @@ let val ex_tm = S.mk_exists{Bvar=v,Body=tm} in (ex_tm, R.CHOOSE(tych v, R.ASSUME (tych ex_tm)) thm) end - val [veq] = filter (can S.dest_eq) (#1 (R.dest_thm thm)) + val [veq] = List.filter (can S.dest_eq) (#1 (R.dest_thm thm)) val {lhs,rhs} = S.dest_eq veq val L = S.free_vars_lr rhs in #2 (U.itlist CHOOSER L (veq,thm)) end; @@ -851,7 +851,7 @@ val (pats,TCsl) = ListPair.unzip pat_TCs_list val case_thm = complete_cases thy pats val domain = (type_of o hd) pats - val Pname = Term.variant (foldr (foldr add_term_names) + val Pname = Term.variant (Library.foldr (Library.foldr add_term_names) (pats::TCsl, [])) "P" val P = Free(Pname, domain --> HOLogic.boolT) val Sinduct = R.SPEC (tych P) Sinduction @@ -862,7 +862,7 @@ val cases = map (fn pat => betapply (Sinduct_assumf, pat)) pats val tasks = U.zip3 cases TCl' (R.CONJUNCTS Rinduct_assum) val proved_cases = map (prove_case fconst thy) tasks - val v = Free (variant (foldr add_term_names (map concl proved_cases, [])) + val v = Free (variant (Library.foldr add_term_names (map concl proved_cases, [])) "v", domain) val vtyped = tych v diff -r 1b3115d1a8df -r 8d8c70b41bab TFL/thry.ML --- a/TFL/thry.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/TFL/thry.ML Thu Mar 03 12:43:01 2005 +0100 @@ -66,12 +66,12 @@ NONE => NONE | SOME {nchotomy, ...} => SOME {nchotomy = nchotomy, - constructors = the (DatatypePackage.constrs_of thy tname)}; + constructors = valOf (DatatypePackage.constrs_of thy tname)}; fun extract_info thy = let val infos = map snd (Symtab.dest (DatatypePackage.get_datatypes thy)) in {case_congs = map (mk_meta_eq o #case_cong) infos, - case_rewrites = flat (map (map mk_meta_eq o #case_rewrites) infos)} + case_rewrites = List.concat (map (map mk_meta_eq o #case_rewrites) infos)} end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/CCL/CCL.ML --- a/src/CCL/CCL.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/CCL/CCL.ML Thu Mar 03 12:43:01 2005 +0100 @@ -59,7 +59,7 @@ fun mk_inj_rl thy rews s = let fun mk_inj_lemmas r = ([arg_cong] RL [(r RS (r RS eq_lemma))]); - val inj_lemmas = flat (map mk_inj_lemmas rews); + val inj_lemmas = List.concat (map mk_inj_lemmas rews); val tac = REPEAT (ares_tac [iffI,allI,conjI] 1 ORELSE eresolve_tac inj_lemmas 1 ORELSE asm_simp_tac (CCL_ss addsimps rews) 1) @@ -99,7 +99,7 @@ fun mk_lemma (ra,rb) = [lemma] RL [ra RS (rb RS eq_lemma)] RL [distinctness RS notE,sym RS (distinctness RS notE)]; in - fun mk_lemmas rls = flat (map mk_lemma (mk_combs pair rls)); + fun mk_lemmas rls = List.concat (map mk_lemma (mk_combs pair rls)); fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs; end; @@ -117,7 +117,7 @@ (fn _ => [simp_tac (CCL_ss addsimps (rls@inj_rls)) 1]) in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end; -fun mkall_dstnct_thms thy defs i_rls xss = flat (map (mk_dstnct_thms thy defs i_rls) xss); +fun mkall_dstnct_thms thy defs i_rls xss = List.concat (map (mk_dstnct_thms thy defs i_rls) xss); (*** Rewriting and Proving ***) diff -r 1b3115d1a8df -r 8d8c70b41bab src/CCL/typecheck.ML --- a/src/CCL/typecheck.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/CCL/typecheck.ML Thu Mar 03 12:43:01 2005 +0100 @@ -85,12 +85,12 @@ fun IHinst tac rls = SUBGOAL (fn (Bi,i) => let val bvs = bvars Bi []; - val ihs = filter could_IH (Logic.strip_assums_hyp Bi); + val ihs = List.filter could_IH (Logic.strip_assums_hyp Bi); val rnames = map (fn x=> let val (a,b) = get_bno [] 0 x - in (nth_elem(a,bvs),b) end) ihs; + in (List.nth(bvs,a),b) end) ihs; fun try_IHs [] = no_tac - | try_IHs ((x,y)::xs) = tac [("g",x)] (nth_elem(y-1,rls)) i ORELSE (try_IHs xs); + | try_IHs ((x,y)::xs) = tac [("g",x)] (List.nth(rls,y-1)) i ORELSE (try_IHs xs); in try_IHs rnames end); (*****) diff -r 1b3115d1a8df -r 8d8c70b41bab src/CTT/CTT.ML --- a/src/CTT/CTT.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/CTT/CTT.ML Thu Mar 03 12:43:01 2005 +0100 @@ -157,7 +157,7 @@ (*0 subgoals vs 1 or more*) val (safe0_brls, safep_brls) = - partition (apl(0,op=) o subgoals_of_brl) safe_brls; + List.partition (apl(0,op=) o subgoals_of_brl) safe_brls; fun safestep_tac thms i = form_tac ORELSE diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOL/IFOL_lemmas.ML --- a/src/FOL/IFOL_lemmas.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOL/IFOL_lemmas.ML Thu Mar 03 12:43:01 2005 +0100 @@ -384,7 +384,7 @@ (*special cases for free variables P, Q, R, S -- up to 3 arguments*) val pred_congs = - flat (map (fn c => + List.concat (map (fn c => map (fn th => read_instantiate [("P",c)] th) [pred1_cong,pred2_cong,pred3_cong]) (explode"PQRS")); diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOL/eqrule_FOL_data.ML --- a/src/FOL/eqrule_FOL_data.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOL/eqrule_FOL_data.ML Thu Mar 03 12:43:01 2005 +0100 @@ -37,14 +37,14 @@ (case Term.head_of p of Const(a,_) => (case Library.assoc(pairs,a) of - SOME(rls) => flat (map atoms ([th] RL rls)) + SOME(rls) => List.concat (map atoms ([th] RL rls)) | NONE => [th]) | _ => [th]) | _ => [th]) in atoms end; val prep_meta_eq = - (mapfilter + (List.mapPartial mk_eq o (mk_atomize tranformation_pairs) o Drule.gen_all diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOL/intprover.ML --- a/src/FOL/intprover.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOL/intprover.ML Thu Mar 03 12:43:01 2005 +0100 @@ -60,7 +60,7 @@ (*0 subgoals vs 1 or more: the p in safep is for positive*) val (safe0_brls, safep_brls) = - partition (apl(0,op=) o subgoals_of_brl) safe_brls; + List.partition (apl(0,op=) o subgoals_of_brl) safe_brls; (*Attack subgoals using safe inferences -- matching, not resolution*) val safe_step_tac = FIRST' [eq_assume_tac, diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOL/simpdata.ML --- a/src/FOL/simpdata.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOL/simpdata.ML Thu Mar 03 12:43:01 2005 +0100 @@ -116,7 +116,7 @@ (case head_of p of Const(a,_) => (case assoc(pairs,a) of - SOME(rls) => flat (map atoms ([th] RL rls)) + SOME(rls) => List.concat (map atoms ([th] RL rls)) | NONE => [th]) | _ => [th]) | _ => [th]) diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOLP/IFOLP.ML --- a/src/FOLP/IFOLP.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOLP/IFOLP.ML Thu Mar 03 12:43:01 2005 +0100 @@ -78,7 +78,7 @@ and concl = discard_proof (Logic.strip_assums_concl prem) in if exists (fn hyp => hyp aconv concl) hyps - then case distinct (filter (fn hyp=> could_unify(hyp,concl)) hyps) of + then case distinct (List.filter (fn hyp=> could_unify(hyp,concl)) hyps) of [_] => assume_tac i | _ => no_tac else no_tac @@ -310,7 +310,7 @@ (*special cases for free variables P, Q, R, S -- up to 3 arguments*) val pred_congs = - flat (map (fn c => + List.concat (map (fn c => map (fn th => read_instantiate [("P",c)] th) [pred1_cong,pred2_cong,pred3_cong]) (explode"PQRS")); diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOLP/classical.ML --- a/src/FOLP/classical.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOLP/classical.ML Thu Mar 03 12:43:01 2005 +0100 @@ -112,7 +112,7 @@ (*Note that allE precedes exI in haz_brls*) fun make_cs {safeIs,safeEs,hazIs,hazEs} = let val (safe0_brls, safep_brls) = (*0 subgoals vs 1 or more*) - partition (apl(0,op=) o subgoals_of_brl) + List.partition (apl(0,op=) o subgoals_of_brl) (sort (make_ord lessb) (joinrules(safeIs, safeEs))) in CS{safeIs=safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=hazEs, safe0_brls=safe0_brls, safep_brls=safep_brls, diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOLP/intprover.ML --- a/src/FOLP/intprover.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOLP/intprover.ML Thu Mar 03 12:43:01 2005 +0100 @@ -50,7 +50,7 @@ (*0 subgoals vs 1 or more: the p in safep is for positive*) val (safe0_brls, safep_brls) = - partition (apl(0,op=) o subgoals_of_brl) safe_brls; + List.partition (apl(0,op=) o subgoals_of_brl) safe_brls; (*Attack subgoals using safe inferences*) val safe_step_tac = FIRST' [uniq_assume_tac, diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOLP/simp.ML --- a/src/FOLP/simp.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOLP/simp.ML Thu Mar 03 12:43:01 2005 +0100 @@ -84,7 +84,7 @@ biresolve_tac (Net.unify_term net (lhs_of (strip_assums_concl prem))) i); -fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); +fun nth_subgoal i thm = List.nth(prems_of thm,i-1); fun goal_concl i thm = strip_assums_concl(nth_subgoal i thm); @@ -161,7 +161,7 @@ in case f of Const(c,T) => if c mem ccs - then foldr add_hvars (args,hvars) + then Library.foldr add_hvars (args,hvars) else add_term_vars(tm,hvars) | _ => add_term_vars(tm,hvars) end @@ -173,7 +173,7 @@ if at then vars else add_term_vars(tm,vars) fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm in if length(tml)=length(al) - then foldr itf (tml~~al,vars) + then Library.foldr itf (tml~~al,vars) else vars end fun add_vars (tm,vars) = case tm of @@ -194,12 +194,12 @@ val lhs = rhs_of_eq 1 thm' val rhs = lhs_of_eq nops thm' val asms = tl(rev(tl(prems_of thm'))) - val hvars = foldr (add_hidden_vars ccs) (lhs::rhs::asms,[]) + val hvars = Library.foldr (add_hidden_vars ccs) (lhs::rhs::asms,[]) val hvars = add_new_asm_vars new_asms (rhs,hvars) fun it_asms (asm,hvars) = if atomic asm then add_new_asm_vars new_asms (asm,hvars) else add_term_frees(asm,hvars) - val hvars = foldr it_asms (asms,hvars) + val hvars = Library.foldr it_asms (asms,hvars) val hvs = map (#1 o dest_Var) hvars val refl1_tac = refl_tac 1 fun norm_step_tac st = st |> @@ -216,7 +216,7 @@ fun add_norm_tags congs = let val ccs = map cong_const congs - val new_asms = filter (exists not o #2) + val new_asms = List.filter (exists not o #2) (ccs ~~ (map (map atomic o prems_of) congs)); in add_norms(congs,ccs,new_asms) end; @@ -252,7 +252,7 @@ (writeln"\nDuplicate rewrite or congruence rule:"; print_thm th; net); -val insert_thms = foldr insert_thm_warn; +val insert_thms = Library.foldr insert_thm_warn; fun addrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thm) = let val thms = mk_simps thm @@ -260,7 +260,7 @@ simps = (thm,thms)::simps, simp_net = insert_thms(thms,simp_net)} end; -val op addrews = foldl addrew; +val op addrews = Library.foldl addrew; fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) = let val congs' = thms @ congs; @@ -278,10 +278,10 @@ (writeln"\nNo such rewrite or congruence rule:"; print_thm th; net); -val delete_thms = foldr delete_thm_warn; +val delete_thms = Library.foldr delete_thm_warn; fun op delcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) = -let val congs' = foldl (gen_rem Drule.eq_thm_prop) (congs,thms) +let val congs' = Library.foldl (gen_rem Drule.eq_thm_prop) (congs,thms) in SS{auto_tac=auto_tac, congs= congs', cong_net= delete_thms(map mk_trans thms,cong_net), mk_simps= normed_rews congs', simps=simps, simp_net=simp_net} @@ -298,7 +298,7 @@ simps = simps', simp_net = delete_thms(thms,simp_net)} end; -val op delrews = foldl delrew; +val op delrews = Library.foldl delrew; fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net,...}, auto_tac) = @@ -426,9 +426,9 @@ SOME(thm',_) => let val ps = prems_of thm and ps' = prems_of thm'; val n = length(ps')-length(ps); - val a = length(strip_assums_hyp(nth_elem(i-1,ps))) + val a = length(strip_assums_hyp(List.nth(ps,i-1))) val l = map (fn p => length(strip_assums_hyp(p))) - (take(n,drop(i-1,ps'))); + (Library.take(n,Library.drop(i-1,ps'))); in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end | NONE => (REW::ss,thm,anet,ats,cs); @@ -437,9 +437,9 @@ fun add_asms(ss,thm,a,anet,ats,cs) = let val As = strip_varify(nth_subgoal i thm, a, []); val thms = map (trivial o cterm_of(#sign(rep_thm(thm))))As; - val new_rws = flat(map mk_rew_rules thms); - val rwrls = map mk_trans (flat(map mk_rew_rules thms)); - val anet' = foldr lhs_insert_thm (rwrls,anet) + val new_rws = List.concat(map mk_rew_rules thms); + val rwrls = map mk_trans (List.concat(map mk_rew_rules thms)); + val anet' = Library.foldr lhs_insert_thm (rwrls,anet) in if !tracing andalso not(null new_rws) then (writeln"Adding rewrites:"; prths new_rws; ()) else (); @@ -539,7 +539,7 @@ fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) = let fun xn_list(x,n) = let val ixs = map (fn i => (x^(radixstring(26,"a",i)),0)) (0 upto n); - in ListPair.map eta_Var (ixs, take(n+1,Ts)) end + in ListPair.map eta_Var (ixs, Library.take(n+1,Ts)) end val lhs = list_comb(f,xn_list("X",k-1)) val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik) in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end; @@ -596,7 +596,7 @@ val T' = incr_tvar 9 T; in mk_cong_type sg (Const(f,T'),T') end; -fun mk_congs thy = flat o map (mk_cong_thy thy); +fun mk_congs thy = List.concat o map (mk_cong_thy thy); fun mk_typed_congs thy = let val sg = sign_of thy; @@ -606,7 +606,7 @@ val t = case Sign.const_type sg f of SOME(_) => Const(f,T) | NONE => Free(f,T) in (t,T) end -in flat o map (mk_cong_type sg o readfT) end; +in List.concat o map (mk_cong_type sg o readfT) end; end (* local *) end (* SIMP *); diff -r 1b3115d1a8df -r 8d8c70b41bab src/FOLP/simpdata.ML --- a/src/FOLP/simpdata.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/FOLP/simpdata.ML Thu Mar 03 12:43:01 2005 +0100 @@ -85,7 +85,7 @@ (case head_of p of Const(a,_) => (case assoc(pairs,a) of - SOME(rls) => flat (map atoms ([th] RL rls)) + SOME(rls) => List.concat (map atoms ([th] RL rls)) | NONE => [th]) | _ => [th]) | _ => [th]) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Algebra/ringsimp.ML --- a/src/HOL/Algebra/ringsimp.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Algebra/ringsimp.ML Thu Mar 03 12:43:01 2005 +0100 @@ -74,17 +74,17 @@ handle TERM _ => []; val prems = ProofContext.prems_of ctxt; - val rings = flat (map (ring_filter o #prop o rep_thm) prems); - val comms = flat (map (comm_filter o #prop o rep_thm) prems); + val rings = List.concat (map (ring_filter o #prop o rep_thm) prems); + val comms = List.concat (map (comm_filter o #prop o rep_thm) prems); val non_comm_rings = rings \\ comms; val comm_rings = rings inter_string comms; val _ = tracing ("Non-commutative rings in proof context: " ^ commas non_comm_rings ^ "\nCommutative rings in proof context: " ^ commas comm_rings); val simps = - flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".ring_simprules", NONE)) + List.concat (map (fn s => ProofContext.get_thms ctxt (s ^ ".ring_simprules", NONE)) non_comm_rings) @ - flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".cring_simprules", NONE)) + List.concat (map (fn s => ProofContext.get_thms ctxt (s ^ ".cring_simprules", NONE)) comm_rings); in Method.SIMPLE_METHOD' HEADGOAL (asm_full_simp_tac (cring_ss addsimps simps)) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Bali/Basis.thy Thu Mar 03 12:43:01 2005 +0100 @@ -188,7 +188,7 @@ (*###to be phased out *) ML {* fun noAll_simpset () = simpset() setmksimps - mksimps (filter (fn (x,_) => x<>"All") mksimps_pairs) + mksimps (List.filter (fn (x,_) => x<>"All") mksimps_pairs) *} lemma All_Ex_refl_eq2 [simp]: diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/HOL.thy --- a/src/HOL/HOL.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/HOL.thy Thu Mar 03 12:43:01 2005 +0100 @@ -771,7 +771,7 @@ fun wrong_prem (Const ("All", _) $ (Abs (_, _, t))) = wrong_prem t | wrong_prem (Bound _) = true | wrong_prem _ = false - val filter_right = filter (fn t => not (wrong_prem (HOLogic.dest_Trueprop (hd (Thm.prems_of t))))) + val filter_right = List.filter (fn t => not (wrong_prem (HOLogic.dest_Trueprop (hd (Thm.prems_of t))))) in fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]) fun smp_tac j = EVERY'[dresolve_tac (smp j), atac] diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Import/hol4rews.ML --- a/src/HOL/Import/hol4rews.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Import/hol4rews.ML Thu Mar 03 12:43:01 2005 +0100 @@ -614,7 +614,7 @@ then let val paths = NameSpace.unpack isa - val i = drop(length paths - 2,paths) + val i = Library.drop(length paths - 2,paths) in case i of [seg,con] => output_thy ^ "." ^ seg ^ "." ^ con diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Import/import_syntax.ML --- a/src/HOL/Import/import_syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Import/import_syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -49,7 +49,7 @@ val thyname = get_generating_thy thy val segname = get_import_segment thy val (int_thms,facts) = Replay.setup_int_thms thyname thy - val thmnames = filter (not o should_ignore thyname thy) facts + val thmnames = List.filter (not o should_ignore thyname thy) facts in thy |> clear_import_thy |> set_segment thyname segname @@ -68,7 +68,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,thmname) => ignore_hol4 thyname thmname thy) (thy,ignored) + Library.foldl (fn (thy,thmname) => ignore_hol4 thyname thmname thy) (thy,ignored) end) val thm_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname) @@ -77,7 +77,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,(hol,isa)) => add_hol4_mapping thyname hol isa thy) (thy,thmmaps) + Library.foldl (fn (thy,(hol,isa)) => add_hol4_mapping thyname hol isa thy) (thy,thmmaps) end) val type_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname) @@ -86,7 +86,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,(hol,isa)) => add_hol4_type_mapping thyname hol false isa thy) (thy,typmaps) + Library.foldl (fn (thy,(hol,isa)) => add_hol4_type_mapping thyname hol false isa thy) (thy,typmaps) end) val def_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname) @@ -95,7 +95,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps) + Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps) end) val const_renames = Scan.repeat1 (P.name --| P.$$$ ">" -- P.name) @@ -104,7 +104,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,(hol,isa)) => add_hol4_const_renaming thyname hol isa thy) (thy,renames) + Library.foldl (fn (thy,(hol,isa)) => add_hol4_const_renaming thyname hol isa thy) (thy,renames) end) val const_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ)) @@ -113,7 +113,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol false isa thy + Library.foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol false isa thy | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol false isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps) end) @@ -123,7 +123,7 @@ let val thyname = get_import_thy thy in - foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol true isa thy + Library.foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol true isa thy | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol true isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps) end) @@ -138,7 +138,7 @@ Scan.empty_lexicon) val get_lexes = fn () => !lexes val token_source = OuterLex.source false get_lexes (Position.line 1) symb_source - val token_list = filter (not o (OuterLex.is_kind OuterLex.Space)) (Source.exhaust token_source) + val token_list = filter_out (OuterLex.is_kind OuterLex.Space) (Source.exhaust token_source) val import_segmentP = OuterParse.$$$ "import_segment" |-- import_segment val type_mapsP = OuterParse.$$$ "type_maps" |-- type_maps val def_mapsP = OuterParse.$$$ "def_maps" |-- def_maps diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Import/proof_kernel.ML --- a/src/HOL/Import/proof_kernel.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Import/proof_kernel.ML Thu Mar 03 12:43:01 2005 +0100 @@ -580,7 +580,7 @@ else find ps end) handle OS.SysErr _ => find ps in - apsome (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path) + Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path) end fun proof_file_name thyname thmname thy = @@ -675,7 +675,7 @@ val (c,asl) = case terms of [] => raise ERR "x2p" "Bad oracle description" | (hd::tl) => (hd,tl) - val tg = foldr (fn (oracle,tg) => Tag.merge (Tag.read oracle) tg) (ors,Tag.empty_tag) + val tg = Library.foldr (fn (oracle,tg) => Tag.merge (Tag.read oracle) tg) (ors,Tag.empty_tag) in mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c)) end @@ -916,7 +916,7 @@ end fun implies_elim_all th = - foldl (fn (th,p) => implies_elim th (assume p)) (th,cprems_of th) + Library.foldl (fn (th,p) => implies_elim th (assume p)) (th,cprems_of th) fun norm_hyps th = th |> beta_eta_thm @@ -1064,7 +1064,7 @@ | process _ = raise ERR "disamb_helper" "Internal error" val (vars',rens',inst) = - foldr process (varstm,(vars,rens,[])) + Library.foldr process (varstm,(vars,rens,[])) in ({vars=vars',rens=rens'},inst) end @@ -1100,7 +1100,7 @@ end fun disamb_terms_from info tms = - foldr (fn (tm,(info,tms)) => + Library.foldr (fn (tm,(info,tms)) => let val (info',tm') = disamb_term_from info tm in @@ -1109,7 +1109,7 @@ (tms,(info,[])) fun disamb_thms_from info hthms = - foldr (fn (hthm,(info,thms)) => + Library.foldr (fn (hthm,(info,thms)) => let val (info',tm') = disamb_thm_from info hthm in @@ -1127,7 +1127,7 @@ let val vars = collect_vars (prop_of th) val (rens',inst,_) = - foldr (fn((ren as (vold as Free(vname,_),vnew)), + Library.foldr (fn((ren as (vold as Free(vname,_),vnew)), (rens,inst,vars)) => (case Library.find_first (fn Free(v,_) => v = vname | _ => false) vars of SOME v' => if v' = vold @@ -1139,7 +1139,7 @@ val (dom,rng) = ListPair.unzip inst val th' = mk_INST (map (cterm_of sg) dom) (map (cterm_of sg) rng) th val nvars = collect_vars (prop_of th') - val rens' = filter (fn(_,v) => v mem nvars) rens + val rens' = List.filter (fn(_,v) => v mem nvars) rens val res = HOLThm(rens',th') in res @@ -1167,7 +1167,7 @@ end fun non_trivial_term_consts tm = - filter (fn c => not (c = "Trueprop" orelse + List.filter (fn c => not (c = "Trueprop" orelse c = "All" orelse c = "op -->" orelse c = "op &" orelse @@ -1239,7 +1239,7 @@ val th1 = (SOME (transform_error (PureThy.get_thm thy) (thmname, NONE)) handle ERROR_MESSAGE _ => (case split_name thmname of - SOME (listname,idx) => (SOME (nth_elem(idx-1,PureThy.get_thms thy (listname, NONE))) + SOME (listname,idx) => (SOME (List.nth(PureThy.get_thms thy (listname, NONE),idx-1)) handle _ => NONE) | NONE => NONE)) in @@ -1794,7 +1794,7 @@ | inst_type ty1 ty2 (ty as Type(name,tys)) = Type(name,map (inst_type ty1 ty2) tys) in - foldr (fn (v,th) => + Library.foldr (fn (v,th) => let val cdom = fst (dom_rng (fst (dom_rng cty))) val vty = type_of v @@ -1806,7 +1806,7 @@ end | SOME _ => raise ERR "GEN_ABS" "Bad constant" | NONE => - foldr (fn (v,th) => mk_ABS v th sg) (vlist',th) + Library.foldr (fn (v,th) => mk_ABS v th sg) (vlist',th) val res = HOLThm(rens_of info',th1) val _ = message "RESULT:" val _ = if_debug pth res @@ -1876,7 +1876,7 @@ let val constname = rename_const thyname thy constname val sg = sign_of thy - val redeclared = is_some (Sign.const_type sg (Sign.intern_const sg constname)); + val redeclared = isSome (Sign.const_type sg (Sign.intern_const sg constname)); val _ = warning ("Introducing constant " ^ constname) val (thmname,thy) = get_defname thyname constname thy val (info,rhs') = disamb_term rhs @@ -1956,21 +1956,21 @@ | dest_exists tm = raise ERR "new_specification" "Bad existential formula" - val (consts,_) = foldl (fn ((cs,ex),cname) => + val (consts,_) = Library.foldl (fn ((cs,ex),cname) => let val (_,cT,p) = dest_exists ex in ((cname,cT,mk_syn thy cname)::cs,p) end) (([],HOLogic.dest_Trueprop (concl_of th)),names) val sg = sign_of thy - val str = foldl (fn (acc,(c,T,csyn)) => + val str = Library.foldl (fn (acc,(c,T,csyn)) => acc ^ "\n " ^ (quotename c) ^ " :: \"" ^ string_of_ctyp (ctyp_of sg T) ^ "\" " ^ (Syntax.string_of_mixfix csyn)) ("consts",consts) val thy' = add_dump str thy in Theory.add_consts_i consts thy' end - val thy1 = foldr (fn(name,thy)=> + val thy1 = Library.foldr (fn(name,thy)=> snd (get_defname thyname name thy)) (names,thy1) fun new_name name = fst (get_defname thyname name thy1) val (thy',res) = SpecificationPackage.add_specification_i NONE @@ -1989,7 +1989,7 @@ then quotename name else (quotename newname) ^ ": " ^ (quotename name),thy') end - val (new_names,thy') = foldr (fn(name,(names,thy)) => + val (new_names,thy') = Library.foldr (fn(name,(names,thy)) => let val (name',thy') = handle_const (name,thy) in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Import/replay.ML --- a/src/HOL/Import/replay.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Import/replay.ML Thu Mar 03 12:43:01 2005 +0100 @@ -25,7 +25,7 @@ end | rp (PSubst(prfs,ctxt,prf)) thy = let - val (thy',ths) = foldr (fn (p,(thy,ths)) => + val (thy',ths) = Library.foldr (fn (p,(thy,ths)) => let val (thy',th) = rp' p thy in @@ -319,7 +319,7 @@ in fst (replay_proof int_thms thyname thmname prf thy) end - val res_thy = foldl replay_fact (thy,thmnames) + val res_thy = Library.foldl replay_fact (thy,thmnames) in res_thy end diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Import/shuffler.ML --- a/src/HOL/Import/shuffler.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Import/shuffler.ML Thu Mar 03 12:43:01 2005 +0100 @@ -42,17 +42,17 @@ case e of THM (msg,i,thms) => (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); - seq print_thm thms) + List.app print_thm thms) | THEORY (msg,thys) => (writeln ("Exception THEORY raised:\n" ^ msg); - seq (Pretty.writeln o Display.pretty_theory) thys) + List.app (Pretty.writeln o Display.pretty_theory) thys) | TERM (msg,ts) => (writeln ("Exception TERM raised:\n" ^ msg); - seq (writeln o Sign.string_of_term sign) ts) + List.app (writeln o Sign.string_of_term sign) ts) | TYPE (msg,Ts,ts) => (writeln ("Exception TYPE raised:\n" ^ msg); - seq (writeln o Sign.string_of_typ sign) Ts; - seq (writeln o Sign.string_of_term sign) ts) + List.app (writeln o Sign.string_of_typ sign) Ts; + List.app (writeln o Sign.string_of_term sign) ts) | e => raise e (*Prints an exception, then fails*) @@ -251,7 +251,7 @@ val tvars = term_tvars t val tfree_names = add_term_tfree_names(t,[]) val (type_inst,_) = - foldl (fn ((inst,used),(w as (v,_),S)) => + Library.foldl (fn ((inst,used),(w as (v,_),S)) => let val v' = variant used v in @@ -553,13 +553,13 @@ end val collect_ignored = - foldr (fn (thm,cs) => + Library.foldr (fn (thm,cs) => let val (lhs,rhs) = Logic.dest_equals (prop_of thm) val ignore_lhs = term_consts lhs \\ term_consts rhs val ignore_rhs = term_consts rhs \\ term_consts lhs in - foldr (op ins_string) (ignore_lhs @ ignore_rhs,cs) + Library.foldr (op ins_string) (ignore_lhs @ ignore_rhs,cs) end) (* set_prop t thms tries to make a theorem with the proposition t from @@ -570,7 +570,7 @@ let val sg = sign_of thy val vars = add_term_frees (t,add_term_vars (t,[])) - val closed_t = foldr (fn (v,body) => let val vT = type_of v + val closed_t = Library.foldr (fn (v,body) => let val vT = type_of v in all vT $ (Abs("x",vT,abstract_over(v,body))) end) (vars,t) val rew_th = norm_term thy closed_t val rhs = snd (dest_equals (cprop_of rew_th)) @@ -614,13 +614,13 @@ val rel_consts = term_consts t \\ ignored val pot_thms = PureThy.thms_containing_consts thy rel_consts in - filter (match_consts ignored t) pot_thms + List.filter (match_consts ignored t) pot_thms end fun gen_shuffle_tac thy search thms i st = let val _ = message ("Shuffling " ^ (string_of_thm st)) - val t = nth_elem(i-1,prems_of st) + val t = List.nth(prems_of st,i-1) val set = set_prop thy t fun process_tac thms st = case set thms of diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Integ/cooper_dec.ML --- a/src/HOL/Integ/cooper_dec.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Integ/cooper_dec.ML Thu Mar 03 12:43:01 2005 +0100 @@ -442,7 +442,7 @@ val ts = coeffs_of t in case ts of [] => raise DVD_UNKNOWN - |_ => foldr (fn(k,r) => r andalso (k mod dn = 0)) (ts,true) + |_ => Library.foldr (fn(k,r) => r andalso (k mod dn = 0)) (ts,true) end; @@ -736,7 +736,7 @@ in (rw,HOLogic.mk_disj(F',rf)) end) val f = if b then linear_add else linear_sub - val p_elements = foldr (fn (i,l) => l union (map (fn e => f [] e (mk_numeral i)) st)) (1 upto d,[]) + val p_elements = Library.foldr (fn (i,l) => l union (map (fn e => f [] e (mk_numeral i)) st)) (1 upto d,[]) in h p_elements end; @@ -873,7 +873,7 @@ fun lift_qelim afn nfn qfn isat = let fun qelim x vars p = let val cjs = conjuncts p - val (ycjs,ncjs) = partition (has_bound) cjs in + val (ycjs,ncjs) = List.partition (has_bound) cjs in (if ycjs = [] then p else let val q = (qfn vars ((HOLogic.exists_const HOLogic.intT ) $ Abs(x,HOLogic.intT,(list_conj ycjs)))) in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Integ/presburger.ML --- a/src/HOL/Integ/presburger.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Integ/presburger.ML Thu Mar 03 12:43:01 2005 +0100 @@ -148,21 +148,21 @@ if body_type T mem [iT, nT] andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] - else foldl op union ([], map getfuncs ts) + else Library.foldl op union ([], map getfuncs ts) | (Var (_, T), ts as _ :: _) => if body_type T mem [iT, nT] andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] - else foldl op union ([], map getfuncs ts) + else Library.foldl op union ([], map getfuncs ts) | (Const (s, T), ts) => if (s, T) mem allowed_consts orelse not (body_type T mem [iT, nT]) - then foldl op union ([], map getfuncs ts) + then Library.foldl op union ([], map getfuncs ts) else [fm] | (Abs (s, T, t), _) => getfuncs t | _ => []; fun abstract_pres sg fm = - foldr (fn (t, u) => + Library.foldr (fn (t, u) => let val T = fastype_of t in all T $ Abs ("x", T, abstract_over (t, u)) end) (getfuncs fm, fm); @@ -198,7 +198,7 @@ abstracted over.*) fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso - map (fn i => snd (nth_elem (i, ps))) (loose_bnos t) @ + map (fn i => snd (List.nth (ps, i))) (loose_bnos t) @ map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_vars t) subset [iT, nT] andalso not (has_free_funcs t); @@ -217,13 +217,13 @@ (HOLogic.all_const T $ Abs (s, T, P), n) else (incr_boundvars ~1 P, n-1) fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; - val (rhs,irhs) = partition (relevant (rev ps)) hs + val (rhs,irhs) = List.partition (relevant (rev ps)) hs val np = length ps - val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) - (ps,(foldr HOLogic.mk_imp (rhs, c), np)) - val (vs, _) = partition (fn t => q orelse (type_of t) = nT) + val (fm',np) = Library.foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) + (ps,(Library.foldr HOLogic.mk_imp (rhs, c), np)) + val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) (term_frees fm' @ term_vars fm'); - val fm2 = foldr mk_all2 (vs, fm') + val fm2 = Library.foldr mk_all2 (vs, fm') in (fm2, np + length vs, length rhs) end; (*Object quantifier to meta --*) @@ -300,7 +300,7 @@ in Method.simple_args (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (foldl op |>) (true, false)) + curry (Library.foldl op |>) (true, false)) (fn (q,a) => fn _ => meth q a 1) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/IsaMakefile Thu Mar 03 12:43:01 2005 +0100 @@ -77,6 +77,7 @@ $(SRC)/Provers/make_elim.ML $(SRC)/Provers/order.ML $(SRC)/Provers/quasi.ML\ $(SRC)/Provers/simplifier.ML $(SRC)/Provers/splitter.ML \ $(SRC)/Provers/trancl.ML \ + $(SRC)/TFL/casesplit.ML \ $(SRC)/TFL/dcterm.ML $(SRC)/TFL/post.ML $(SRC)/TFL/rules.ML \ $(SRC)/TFL/tfl.ML $(SRC)/TFL/thms.ML $(SRC)/TFL/thry.ML \ $(SRC)/TFL/usyntax.ML $(SRC)/TFL/utils.ML \ diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Lambda/WeakNorm.thy --- a/src/HOL/Lambda/WeakNorm.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Lambda/WeakNorm.thy Thu Mar 03 12:43:01 2005 +0100 @@ -566,7 +566,7 @@ | term_of_dB' _ _ = error "term_of_dB: term not in normal form"; fun typing_of_term Ts e (Bound i) = - rtypingT_Var (e, nat_of_int i, dBtype_of_typ (nth_elem (i, Ts))) + rtypingT_Var (e, nat_of_int i, dBtype_of_typ (List.nth (Ts, i))) | typing_of_term Ts e (t $ u) = (case fastype_of1 (Ts, t) of Type ("fun", [T, U]) => rtypingT_App (e, dB_of_term t, dBtype_of_typ T, dBtype_of_typ U, dB_of_term u, diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Library/EfficientNat.thy --- a/src/HOL/Library/EfficientNat.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Library/EfficientNat.thy Thu Mar 03 12:43:01 2005 +0100 @@ -100,7 +100,7 @@ val Suc_if_eq' = Thm.transfer thy Suc_if_eq; val ctzero = cterm_of (sign_of Main.thy) HOLogic.zero; val vname = variant (map fst - (foldl add_term_varnames ([], map prop_of thms))) "x"; + (Library.foldl add_term_varnames ([], map prop_of thms))) "x"; val cv = cterm_of (sign_of Main.thy) (Var ((vname, 0), HOLogic.natT)); fun lhs_of th = snd (Thm.dest_comb (fst (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th)))))); @@ -114,13 +114,13 @@ map (apfst (pairself (Thm.capply ct1))) (find_vars ct2) end | _ => []); - val eqs1 = flat (map + val eqs1 = List.concat (map (fn th => map (pair th) (find_vars (lhs_of th))) thms); - val eqs2 = map (fn x as (_, ((_, ctzero), _)) => (x, mapfilter (fn th => + val eqs2 = map (fn x as (_, ((_, ctzero), _)) => (x, List.mapPartial (fn th => SOME (th, Thm.cterm_match (lhs_of th, ctzero)) handle Pattern.MATCH => NONE) thms)) eqs1; fun distinct_vars vs = forall is_Var vs andalso null (duplicates vs); - val eqs2' = map (apsnd (filter (fn (_, (_, s)) => + val eqs2' = map (apsnd (List.filter (fn (_, (_, s)) => distinct_vars (map (term_of o snd) s)))) eqs2; fun mk_thms b ((th, ((ct, _), cv')), (th', s) :: _) = let @@ -133,7 +133,7 @@ SOME (rhs_of th''), SOME (Thm.cabs cv' (rhs_of th)), SOME cv'] Suc_if_eq')) th'') (Thm.forall_intr cv' th) in - mapfilter (fn thm => + List.mapPartial (fn thm => if thm = th then SOME th''' else if b andalso thm = th' then NONE else SOME thm) thms @@ -160,13 +160,13 @@ let val Suc_clause' = Thm.transfer thy Suc_clause; val vname = variant (map fst - (foldl add_term_varnames ([], map prop_of thms))) "x"; + (Library.foldl add_term_varnames ([], map prop_of thms))) "x"; fun find_var (t as Const ("Suc", _) $ (v as Var _)) = SOME (t, v) | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x) | find_var _ = NONE; fun find_thm th = let val th' = ObjectLogic.atomize_thm th - in apsome (pair (th, th')) (find_var (prop_of th')) end + in Option.map (pair (th, th')) (find_var (prop_of th')) end in case get_first find_thm thms of NONE => thms @@ -191,8 +191,8 @@ let val dest = fst o HOLogic.dest_mem o HOLogic.dest_Trueprop in if forall (can (dest o concl_of)) ths andalso - exists (fn th => "Suc" mem foldr add_term_consts - (mapfilter (try dest) (concl_of th :: prems_of th), [])) ths + exists (fn th => "Suc" mem Library.foldr add_term_consts + (List.mapPartial (try dest) (concl_of th :: prems_of th), [])) ths then remove_suc_clause thy ths else ths end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/List.thy --- a/src/HOL/List.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/List.thy Thu Mar 03 12:43:01 2005 +0100 @@ -2127,7 +2127,7 @@ (let val c = chr (dest_nibble c1 * 16 + dest_nibble c2) in if Symbol.is_printable c then SOME (gr, Pretty.quote (Pretty.str c)) else NONE - end handle LIST _ => NONE | Match => NONE) + end handle Fail _ => NONE | Match => NONE) | char_codegen thy gr dep b _ = NONE; in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Matrix/Cplex.ML --- a/src/HOL/Matrix/Cplex.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Matrix/Cplex.ML Thu Mar 03 12:43:01 2005 +0100 @@ -716,7 +716,7 @@ fun merge a b = Symtab.merge (op =) (a, b) -fun mergemap f ts = foldl +fun mergemap f ts = Library.foldl (fn (table, x) => merge table (f x)) (Symtab.empty, ts) @@ -778,7 +778,7 @@ (fn (l, (k,v)) => (make_pos_constr k) :: l) ([], positive_vars)) val bound_constrs = map (fn x => (NONE, x)) - (flat (map bound2constr bounds)) + (List.concat (map bound2constr bounds)) val constraints' = constraints @ pos_constrs @ bound_constrs val bounds' = rev(Symtab.foldl (fn (l, (v,_)) => (make_free_bound v)::l) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Matrix/CplexMatrixConverter.ML --- a/src/HOL/Matrix/CplexMatrixConverter.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Matrix/CplexMatrixConverter.ML Thu Mar 03 12:43:01 2005 +0100 @@ -73,7 +73,7 @@ set_elem vec (s2i v) (if positive then num else "-"^num) | setprod _ _ _ = raise (Converter "term is not a normed product") - fun sum2vec (cplexSum ts) = foldl (fn (vec, t) => setprod vec true t) (empty_vector, ts) + fun sum2vec (cplexSum ts) = Library.foldl (fn (vec, t) => setprod vec true t) (empty_vector, ts) | sum2vec t = setprod empty_vector true t fun constrs2Ab j A b [] = (A, b) @@ -100,7 +100,7 @@ let fun setv (v, (name, value)) = (matrix_builder.set_elem v (name2index name) value) in - (opt, foldl setv (matrix_builder.empty_vector, entries)) + (opt, Library.foldl setv (matrix_builder.empty_vector, entries)) end | convert_results _ _ = raise (Converter "No optimal result") diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Matrix/codegen_prep.ML --- a/src/HOL/Matrix/codegen_prep.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Matrix/codegen_prep.ML Thu Mar 03 12:43:01 2005 +0100 @@ -80,7 +80,7 @@ fun prepare_thms ths = let - val ths = (filter is_meta_eq ths) @ (map (standard o mk_meta_eq) (filter (not o is_meta_eq) ths)) + val ths = (List.filter is_meta_eq ths) @ (map (standard o mk_meta_eq) (filter_out is_meta_eq ths)) val _ = if forall Thm.no_prems ths then () else raise (Prepare_thms "premisse found") val thmgroups = group head_name ths val test_clashgroups = group (fn (a,b) => mangle_id a) thmgroups diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Matrix/eq_codegen.ML --- a/src/HOL/Matrix/eq_codegen.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Matrix/eq_codegen.ML Thu Mar 03 12:43:01 2005 +0100 @@ -24,7 +24,7 @@ fun parens b = if b then Pretty.enclose "(" ")" else Pretty.block; fun gen_mk_val f xs ps = Pretty.block ([Pretty.str "val ", - f (length xs > 1) (flat + f (length xs > 1) (List.concat (separate [Pretty.str ",", Pretty.brk 1] (map (single o Pretty.str) xs))), Pretty.str " =", Pretty.brk 1] @ ps @ [Pretty.str ";"]); @@ -79,7 +79,7 @@ let val vs' = variantlist (map mk_decomp_tname Ts, vs); val (vs'', bs', ps') = - foldl decomp_type_code ((vs @ vs', bs, ps @ + Library.foldl decomp_type_code ((vs @ vs', bs, ps @ [mk_vall vs' [Pretty.str "Thm.dest_ctyp", Pretty.brk 1, Pretty.str v]]), vs' ~~ Ts) in @@ -106,7 +106,7 @@ fun pretty_pattern b (Const (s, _)) = Pretty.block [Pretty.str "Const", Pretty.brk 1, Pretty.str ("(\"" ^ s ^ "\", _)")] | pretty_pattern b (t as _ $ _) = parens b - (flat (separate [Pretty.str " $", Pretty.brk 1] + (List.concat (separate [Pretty.str " $", Pretty.brk 1] (map (single o pretty_pattern true) (op :: (strip_comb t))))) | pretty_pattern b _ = Pretty.str "_"; @@ -129,7 +129,7 @@ [] => Pretty.str s | Ts => parens b [Pretty.str "tyinst_cterm", Pretty.brk 1, Pretty.list "[" "]" (map (fn (ixn, _) => mk_tyinst - (ixn, the (assoc (ty_bs, ixn)))) Ts), + (ixn, valOf (assoc (ty_bs, ixn)))) Ts), Pretty.brk 1, Pretty.str s]); fun mk_cterm_code b ty_bs ts xs (vals, t $ u) = @@ -151,7 +151,7 @@ in (vals', parens b [Pretty.str "Thm.cabs", Pretty.brk 1, p, Pretty.brk 1, p']) end - | mk_cterm_code b ty_bs ts xs (vals, Bound i) = (vals, nth_elem (i, xs)) + | mk_cterm_code b ty_bs ts xs (vals, Bound i) = (vals, List.nth (xs, i)) | mk_cterm_code b ty_bs ts xs (vals, t) = (case assoc (vals, t) of NONE => let val SOME s = assoc (ts, t) @@ -182,7 +182,7 @@ in if s = s' then (vs, subs) else (s' :: vs, ((s, i), Var ((s', i), T)) :: subs) end; - val (vs', subs) = foldl mk_subst ((vs, []), term_vars t) + val (vs', subs) = Library.foldl mk_subst ((vs, []), term_vars t) in (vs', subst_Vars subs t) end; fun is_instance sg t u = t = subst_TVars_Vartab @@ -190,7 +190,7 @@ (fastype_of u, fastype_of t))) u handle Type.TYPE_MATCH => false; (* -fun lookup sg fs t = apsome snd (Library.find_first +fun lookup sg fs t = Option.map snd (Library.find_first (is_instance sg t o fst) fs); *) @@ -269,18 +269,18 @@ val cvs' = map Var (map (rpair 0) ys ~~ Ts); val rs = cvs' ~~ cvs; val lhs = list_comb (Const (cname, Ts ---> uT), cvs'); - val rhs = foldl betapply (f, cvs'); + val rhs = Library.foldl betapply (f, cvs'); val (vs', tm_bs, tm_vals) = decomp_term_code false ((vs @ ys, [], []), (ct, lhs)); val ((vs'', all_vals), (eq', ct')) = mk_simpl_code sg case_tab false fs (tm_bs @ ts) ty_bs thm_bs ((vs', vals), rhs); val (old_vals, eq_vals) = splitAt (i, all_vals); - val vs''' = vs @ filter (fn v => exists + val vs''' = vs @ List.filter (fn v => exists (fn (_, ((v', _), _)) => v = v') old_vals) (vs'' \\ vs'); val insts2 = map (fn (t, s) => Pretty.block [Pretty.str "(", - inst_ty false ty_bs' t (the (assoc (thm_bs, t))), Pretty.str ",", + inst_ty false ty_bs' t (valOf (assoc (thm_bs, t))), Pretty.str ",", Pretty.brk 1, Pretty.str (s ^ ")")]) ((fvs ~~ xs) @ - (map (fn (v, s) => (the (assoc (rs, v)), s)) tm_bs)); + (map (fn (v, s) => (valOf (assoc (rs, v)), s)) tm_bs)); val eq'' = if null insts1 andalso null insts2 then Pretty.str name else parens (eq' <> "") [Pretty.str (if null cvs then "Thm.instantiate" else "Drule.instantiate"), @@ -293,7 +293,7 @@ in ((vs''', old_vals), Pretty.block [pretty_pattern false lhs, Pretty.str " =>", - Pretty.brk 1, mk_let "let" 2 (tm_vals @ flat (map (snd o snd) eq_vals)) + Pretty.brk 1, mk_let "let" 2 (tm_vals @ List.concat (map (snd o snd) eq_vals)) [Pretty.str ("(" ^ ct' ^ ","), Pretty.brk 1, eq''', Pretty.str ")"]]) end; @@ -308,7 +308,7 @@ fvals @ ty_vals @ [mk_val [ct', eq'] ([Pretty.str "(case", Pretty.brk 1, Pretty.str ("term_of " ^ ct ^ " of"), Pretty.brk 1] @ - flat (separate [Pretty.brk 1, Pretty.str "| "] + List.concat (separate [Pretty.brk 1, Pretty.str "| "] (map single case_ps)) @ [Pretty.str ")"])] in if b then @@ -334,7 +334,7 @@ val b = forall (unint sg fs) us; val (q, eqs) = foldl_map (mk_simpl_code sg case_tab (not b) fs ts ty_bs thm_bs) ((vs, vals), us); - val ((vs', vals'), (eqf, ctf)) = if is_some (lookup sg fs f) andalso b + val ((vs', vals'), (eqf, ctf)) = if isSome (lookup sg fs f) andalso b then (q, ("", "")) else mk_simpl_code sg case_tab (not b) fs ts ty_bs thm_bs (q, f); val ct = variant vs' "ct"; @@ -372,38 +372,38 @@ fun mk_funs_code sg case_tab fs fs' = let - val case_thms = mapfilter (fn s => (case Symtab.lookup (case_tab, s) of + val case_thms = List.mapPartial (fn s => (case Symtab.lookup (case_tab, s) of NONE => NONE | SOME thms => SOME (unsuffix "_case" (Sign.base_name s) ^ ".cases", map (fn i => Sign.base_name s ^ "_" ^ string_of_int i) (1 upto length thms) ~~ thms))) - (foldr add_term_consts (map (prop_of o snd) - (flat (map (#3 o snd) fs')), [])); + (Library.foldr add_term_consts (map (prop_of o snd) + (List.concat (map (#3 o snd) fs')), [])); val case_vals = map (fn (s, cs) => mk_vall (map fst cs) [Pretty.str "map my_mk_meta_eq", Pretty.brk 1, Pretty.str ("(thms \"" ^ s ^ "\")")]) case_thms; - val (vs, thm_bs, thm_vals) = foldl mk_term_bindings (([], [], []), - flat (map (map (apsnd prop_of) o #3 o snd) fs') @ - map (apsnd prop_of) (flat (map snd case_thms))); + val (vs, thm_bs, thm_vals) = Library.foldl mk_term_bindings (([], [], []), + List.concat (map (map (apsnd prop_of) o #3 o snd) fs') @ + map (apsnd prop_of) (List.concat (map snd case_thms))); fun mk_fun_code (prfx, (fname, d, eqns)) = let val (f, ts) = strip_comb (lhs_of (snd (hd eqns))); val args = variantlist (replicate (length ts) "ct", vs); - val (vs', ty_bs, ty_vals) = foldl mk_type_bindings + val (vs', ty_bs, ty_vals) = Library.foldl mk_type_bindings ((vs @ args, [], []), args ~~ map fastype_of ts); val insts1 = map mk_tyinst ty_bs; fun mk_eqn_code (name, eqn) = let val (_, argts) = strip_comb (lhs_of eqn); - val (vs'', tm_bs, tm_vals) = foldl (decomp_term_code false) + val (vs'', tm_bs, tm_vals) = Library.foldl (decomp_term_code false) ((vs', [], []), args ~~ argts); val ((vs''', eq_vals), (eq, ct)) = mk_simpl_code sg case_tab false fs (tm_bs @ filter_out (is_Var o fst) thm_bs) ty_bs thm_bs ((vs'', []), rhs_of eqn); val insts2 = map (fn (t, s) => Pretty.block [Pretty.str "(", - inst_ty false ty_bs t (the (assoc (thm_bs, t))), Pretty.str ",", Pretty.brk 1, + inst_ty false ty_bs t (valOf (assoc (thm_bs, t))), Pretty.str ",", Pretty.brk 1, Pretty.str (s ^ ")")]) tm_bs val eq' = if null insts1 andalso null insts2 then Pretty.str name else parens (eq <> "") [Pretty.str "Thm.instantiate", @@ -417,7 +417,7 @@ Pretty.block [parens (length argts > 1) (Pretty.commas (map (pretty_pattern false) argts)), Pretty.str " =>", - Pretty.brk 1, mk_let "let" 2 (ty_vals @ tm_vals @ flat (map (snd o snd) eq_vals)) + Pretty.brk 1, mk_let "let" 2 (ty_vals @ tm_vals @ List.concat (map (snd o snd) eq_vals)) [Pretty.str ("(" ^ ct ^ ","), Pretty.brk 1, eq'', Pretty.str ")"]] end; @@ -441,7 +441,7 @@ [Pretty.str " =", Pretty.brk 1, Pretty.str "(case", Pretty.brk 1, Pretty.list "(" ")" (map (fn s => Pretty.str ("term_of " ^ s)) args), Pretty.str " of", Pretty.brk 1] @ - flat (separate [Pretty.brk 1, Pretty.str "| "] + List.concat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o mk_eqn_code) eqns)) @ [Pretty.str ")"] @ default)) end; @@ -458,19 +458,19 @@ val eqns' = map attach_term eqns; fun mk_node (s, _, (_, th) :: _) = (s, get_head th); fun mk_edges (s, _, ths) = map (pair s) (distinct - (mapfilter (fn t => apsome #1 (lookup sg eqns' t)) - (flat (map (term_consts' o prop_of o snd) ths)))); - val gr = foldr (uncurry Graph.add_edge) - (map (pair "" o #1) eqns @ flat (map mk_edges eqns), - foldr (uncurry Graph.new_node) + (List.mapPartial (fn t => Option.map #1 (lookup sg eqns' t)) + (List.concat (map (term_consts' o prop_of o snd) ths)))); + val gr = Library.foldr (uncurry Graph.add_edge) + (map (pair "" o #1) eqns @ List.concat (map mk_edges eqns), + Library.foldr (uncurry Graph.new_node) (("", Bound 0) :: map mk_node eqns, Graph.empty)); val keys = rev (Graph.all_succs gr [""] \ ""); fun gr_ord (x :: _, y :: _) = int_ord (find_index (equal x) keys, find_index (equal y) keys); - val scc = map (fn xs => filter (fn (_, (s, _, _)) => s mem xs) eqns') + val scc = map (fn xs => List.filter (fn (_, (s, _, _)) => s mem xs) eqns') (sort gr_ord (Graph.strong_conn gr \ [""])); in - flat (separate [Pretty.str ";", Pretty.fbrk, Pretty.str " ", Pretty.fbrk] + List.concat (separate [Pretty.str ";", Pretty.fbrk, Pretty.str " ", Pretty.fbrk] (map (fn eqns'' => [mk_funs_code sg case_tab eqns' eqns'']) scc)) @ [Pretty.str ";", Pretty.fbrk] end; @@ -482,7 +482,7 @@ let val (i', eqs') = foldl_map attach_name (i, eqs) in (i', (s, b, eqs')) end; val (_, eqns') = foldl_map attach_names (1, eqns); - val (names, thms) = split_list (flat (map #3 eqns')); + val (names, thms) = split_list (List.concat (map #3 eqns')); val s = setmp print_mode [] Pretty.string_of (mk_let "local" 2 [mk_vall names [Pretty.str "!SimprocsCodegen.simp_thms"]] (mk_simprocs_code sg eqns')) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Matrix/fspmlp.ML --- a/src/HOL/Matrix/fspmlp.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Matrix/fspmlp.ML Thu Mar 03 12:43:01 2005 +0100 @@ -145,7 +145,7 @@ NONE => NONE | SOME src_sure_bound => SOME (FloatArith.add x (mult_btype src_sure_bound coeff))) in - case foldl add_src_bound (SOME row_bound, sources) of + case Library.foldl add_src_bound (SOME row_bound, sources) of NONE => sure_bound | new_sure_bound as (SOME new_bound) => (case sure_bound of @@ -234,7 +234,7 @@ add_edge g (src_index, LOWER) dest_key row_index coeff end in - foldl fold_src_nodes ((add_row_bound g dest_key row_index row_bound), approx_a) + Library.foldl fold_src_nodes ((add_row_bound g dest_key row_index row_bound), approx_a) end end in @@ -251,7 +251,7 @@ else g end - | _ => foldl fold_dest_nodes (g, approx_a) + | _ => Library.foldl fold_dest_nodes (g, approx_a) end val g = FloatSparseMatrixBuilder.m_fold calcr VarGraph.empty A diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/MicroJava/BV/BVExample.thy Thu Mar 03 12:43:01 2005 +0100 @@ -454,7 +454,7 @@ "op :" ("(_ mem _)") "op Un" ("(_ union _)") "image" ("map") - "UNION" ("(fn A => fn f => flat (map f A))") + "UNION" ("(fn A => fn f => List.concat (map f A))") "Bex" ("(fn A => fn f => exists f A)") "Ball" ("(fn A => fn f => forall f A)") "some_elem" ("hd") diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Modelcheck/EindhovenSyn.ML --- a/src/HOL/Modelcheck/EindhovenSyn.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Modelcheck/EindhovenSyn.ML Thu Mar 03 12:43:01 2005 +0100 @@ -7,7 +7,7 @@ fun mc_eindhoven_tac i state = let val sign = #sign (rep_thm state) in -case drop(i-1,prems_of state) of +case Library.drop(i-1,prems_of state) of [] => Seq.empty | subgoal::_ => let val concl = Logic.strip_imp_concl subgoal; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Modelcheck/MuckeSyn.ML --- a/src/HOL/Modelcheck/MuckeSyn.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Modelcheck/MuckeSyn.ML Thu Mar 03 12:43:01 2005 +0100 @@ -51,14 +51,14 @@ REPEAT (resolve_tac prems 1)]); val sig_move_thm = #sign (rep_thm move_thm); - val bCterm = cterm_of sig_move_thm (the (search_var "b" (concl_of move_thm))); - val aCterm = cterm_of sig_move_thm (the (search_var "a" (hd(prems_of move_thm)))); + val bCterm = cterm_of sig_move_thm (valOf (search_var "b" (concl_of move_thm))); + val aCterm = cterm_of sig_move_thm (valOf (search_var "a" (hd(prems_of move_thm)))); in fun move_mus i state = let val sign = #sign (rep_thm state); - val (subgoal::_) = drop(i-1,prems_of state); + val (subgoal::_) = Library.drop(i-1,prems_of state); val concl = Logic.strip_imp_concl subgoal; (* recursive mu's in prems? *) val redex = search_mu concl; val idx = let val t = #maxidx (rep_thm state) in @@ -84,7 +84,7 @@ (* Normally t can be user-instantiated by the value thy of the Isabelle context *) fun call_mucke_tac i state = let val sign = #sign (rep_thm state); - val (subgoal::_) = drop(i-1,prems_of state); + val (subgoal::_) = Library.drop(i-1,prems_of state); val OraAss = invoke_oracle MuckeSyn.thy "Mucke" (sign,MuckeOracleExn (subgoal)); in (cut_facts_tac [OraAss] i) state @@ -166,7 +166,7 @@ (* the interface *) fun mc_mucke_tac defs i state = - (case drop (i - 1, Thm.prems_of state) of + (case Library.drop (i - 1, Thm.prems_of state) of [] => PureGeneral.Seq.empty | subgoal :: _ => EVERY [ diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Modelcheck/mucke_oracle.ML --- a/src/HOL/Modelcheck/mucke_oracle.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Modelcheck/mucke_oracle.ML Thu Mar 03 12:43:01 2005 +0100 @@ -96,7 +96,7 @@ fun if_full_simp_tac sset i state = let val sign = #sign (rep_thm state); - val (subgoal::_) = drop(i-1,prems_of state); + val (subgoal::_) = Library.drop(i-1,prems_of state); val prems = Logic.strip_imp_prems subgoal; val concl = Logic.strip_imp_concl subgoal; val prems = prems @ [concl]; @@ -253,10 +253,10 @@ fun newName (Var(ix,_), (pairs,used)) = let val v = variant used (string_of_indexname ix) in ((ix,v)::pairs, v::used) end; - val (alist, _) = foldr newName (vars, ([], used)); + val (alist, _) = Library.foldr newName (vars, ([], used)); fun mk_inst (Var(v,T)) = (Var(v,T), - Free(the (assoc(alist,v)), T)); + Free(valOf (assoc(alist,v)), T)); val insts = map mk_inst vars; in subst_atomic insts t end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Product_Type.thy Thu Mar 03 12:43:01 2005 +0100 @@ -816,7 +816,7 @@ val (gr1, qs) = foldl_map mk_code (gr, ps); val (gr2, pu) = Codegen.invoke_codegen thy dep false (gr1, u) in - SOME (gr2, Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, flat + SOME (gr2, Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, List.concat (separate [Pretty.str ";", Pretty.brk 1] (map (fn (pl, pr) => [Pretty.block [Pretty.str "val ", pl, Pretty.str " =", Pretty.brk 1, pr]]) qs))), diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Prolog/HOHH.ML --- a/src/HOL/Prolog/HOHH.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Prolog/HOHH.ML Thu Mar 03 12:43:01 2005 +0100 @@ -63,7 +63,7 @@ imp_all]; (* "((!x. D) :- G) = (!x. D :- G)" *) (*val hyp_resolve_tac = METAHYPS (fn prems => - resolve_tac (flat (map atomizeD prems)) 1); + resolve_tac (List.concat (map atomizeD prems)) 1); -- is nice, but cannot instantiate unknowns in the assumptions *) fun hyp_resolve_tac i st = let fun ap (Const("All",_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a ,t)) @@ -98,10 +98,10 @@ else no_tac); val ptacs = mapn (fn n => fn t => pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems; - in foldl (op APPEND) (no_tac, ptacs) st end; + in Library.foldl (op APPEND) (no_tac, ptacs) st end; fun ptac prog = let - val proga = flat (map atomizeD prog) (* atomize the prog *) + val proga = List.concat (map atomizeD prog) (* atomize the prog *) in (REPEAT_DETERM1 o FIRST' [ rtac TrueI, (* "True" *) rtac conjI, (* "[| P; Q |] ==> P & Q" *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/Presburger/cooper_dec.ML --- a/src/HOL/Tools/Presburger/cooper_dec.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/Presburger/cooper_dec.ML Thu Mar 03 12:43:01 2005 +0100 @@ -442,7 +442,7 @@ val ts = coeffs_of t in case ts of [] => raise DVD_UNKNOWN - |_ => foldr (fn(k,r) => r andalso (k mod dn = 0)) (ts,true) + |_ => Library.foldr (fn(k,r) => r andalso (k mod dn = 0)) (ts,true) end; @@ -736,7 +736,7 @@ in (rw,HOLogic.mk_disj(F',rf)) end) val f = if b then linear_add else linear_sub - val p_elements = foldr (fn (i,l) => l union (map (fn e => f [] e (mk_numeral i)) st)) (1 upto d,[]) + val p_elements = Library.foldr (fn (i,l) => l union (map (fn e => f [] e (mk_numeral i)) st)) (1 upto d,[]) in h p_elements end; @@ -873,7 +873,7 @@ fun lift_qelim afn nfn qfn isat = let fun qelim x vars p = let val cjs = conjuncts p - val (ycjs,ncjs) = partition (has_bound) cjs in + val (ycjs,ncjs) = List.partition (has_bound) cjs in (if ycjs = [] then p else let val q = (qfn vars ((HOLogic.exists_const HOLogic.intT ) $ Abs(x,HOLogic.intT,(list_conj ycjs)))) in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/Presburger/presburger.ML --- a/src/HOL/Tools/Presburger/presburger.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/Presburger/presburger.ML Thu Mar 03 12:43:01 2005 +0100 @@ -148,21 +148,21 @@ if body_type T mem [iT, nT] andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] - else foldl op union ([], map getfuncs ts) + else Library.foldl op union ([], map getfuncs ts) | (Var (_, T), ts as _ :: _) => if body_type T mem [iT, nT] andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] - else foldl op union ([], map getfuncs ts) + else Library.foldl op union ([], map getfuncs ts) | (Const (s, T), ts) => if (s, T) mem allowed_consts orelse not (body_type T mem [iT, nT]) - then foldl op union ([], map getfuncs ts) + then Library.foldl op union ([], map getfuncs ts) else [fm] | (Abs (s, T, t), _) => getfuncs t | _ => []; fun abstract_pres sg fm = - foldr (fn (t, u) => + Library.foldr (fn (t, u) => let val T = fastype_of t in all T $ Abs ("x", T, abstract_over (t, u)) end) (getfuncs fm, fm); @@ -198,7 +198,7 @@ abstracted over.*) fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso - map (fn i => snd (nth_elem (i, ps))) (loose_bnos t) @ + map (fn i => snd (List.nth (ps, i))) (loose_bnos t) @ map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_vars t) subset [iT, nT] andalso not (has_free_funcs t); @@ -217,13 +217,13 @@ (HOLogic.all_const T $ Abs (s, T, P), n) else (incr_boundvars ~1 P, n-1) fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; - val (rhs,irhs) = partition (relevant (rev ps)) hs + val (rhs,irhs) = List.partition (relevant (rev ps)) hs val np = length ps - val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) - (ps,(foldr HOLogic.mk_imp (rhs, c), np)) - val (vs, _) = partition (fn t => q orelse (type_of t) = nT) + val (fm',np) = Library.foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) + (ps,(Library.foldr HOLogic.mk_imp (rhs, c), np)) + val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) (term_frees fm' @ term_vars fm'); - val fm2 = foldr mk_all2 (vs, fm') + val fm2 = Library.foldr mk_all2 (vs, fm') in (fm2, np + length vs, length rhs) end; (*Object quantifier to meta --*) @@ -300,7 +300,7 @@ in Method.simple_args (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (foldl op |>) (true, false)) + curry (Library.foldl op |>) (true, false)) (fn (q,a) => fn _ => meth q a 1) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_abs_proofs.ML --- a/src/HOL/Tools/datatype_abs_proofs.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_abs_proofs.ML Thu Mar 03 12:43:01 2005 +0100 @@ -54,9 +54,9 @@ let val _ = message "Proving case distinction theorems ..."; - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val newTs = take (length (hd descr), recTs); + val newTs = Library.take (length (hd descr), recTs); val {maxidx, ...} = rep_thm induct; val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); @@ -67,11 +67,11 @@ Abs ("z", T', Const ("True", T''))) induct_Ps; val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $ Var (("P", 0), HOLogic.boolT)) - val insts = take (i, dummyPs) @ (P::(drop (i + 1, dummyPs))); + val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs))); val cert = cterm_of (Theory.sign_of thy); val insts' = (map cert induct_Ps) ~~ (map cert insts); - val induct' = refl RS ((nth_elem (i, - split_conj_thm (cterm_instantiate insts' induct))) RSN (2, rev_mp)) + val induct' = refl RS ((List.nth + (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp)) in prove_goalw_cterm [] (cert t) (fn prems => [rtac induct' 1, @@ -95,10 +95,10 @@ val big_name = space_implode "_" new_type_names; val thy0 = add_path flat_names big_name thy; - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val used = foldr add_typ_tfree_names (recTs, []); - val newTs = take (length (hd descr), recTs); + val used = Library.foldr add_typ_tfree_names (recTs, []); + val newTs = Library.take (length (hd descr), recTs); val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); @@ -127,27 +127,27 @@ in (case (strip_dtyp dt, strip_type U) of ((_, DtRec m), (Us, _)) => let - val free2 = mk_Free "y" (Us ---> nth_elem (m, rec_result_Ts)) k; + val free2 = mk_Free "y" (Us ---> List.nth (rec_result_Ts, m)) k; val i = length Us in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all (map (pair "x") Us, HOLogic.mk_mem (HOLogic.mk_prod (app_bnds free1 i, app_bnds free2 i), - nth_elem (m, rec_sets)))) :: prems, + List.nth (rec_sets, m)))) :: prems, free1::t1s, free2::t2s) end | _ => (j + 1, k, prems, free1::t1s, t2s)) end; val Ts = map (typ_of_dtyp descr' sorts) cargs; - val (_, _, prems, t1s, t2s) = foldr mk_prem (cargs ~~ Ts, (1, 1, [], [], [])) + val (_, _, prems, t1s, t2s) = Library.foldr mk_prem (cargs ~~ Ts, (1, 1, [], [], [])) in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop (HOLogic.mk_mem (HOLogic.mk_prod (list_comb (Const (cname, Ts ---> T), t1s), - list_comb (nth_elem (l, rec_fns), t1s @ t2s)), set_name)))], l + 1) + list_comb (List.nth (rec_fns, l), t1s @ t2s)), set_name)))], l + 1) end; - val (rec_intr_ts, _) = foldl (fn (x, ((d, T), set_name)) => - foldl (make_rec_intr T set_name) (x, #3 (snd d))) + val (rec_intr_ts, _) = Library.foldl (fn (x, ((d, T), set_name)) => + Library.foldl (make_rec_intr T set_name) (x, #3 (snd d))) (([], 0), descr' ~~ recTs ~~ rec_sets); val (thy1, {intrs = rec_intrs, elims = rec_elims, ...}) = @@ -163,16 +163,16 @@ let val distinct_tac = (etac Pair_inject 1) THEN (if i < length newTs then - full_simp_tac (HOL_ss addsimps (nth_elem (i, dist_rewrites))) 1 + full_simp_tac (HOL_ss addsimps (List.nth (dist_rewrites, i))) 1 else full_simp_tac dist_ss 1); val inject = map (fn r => r RS iffD1) - (if i < length newTs then nth_elem (i, constr_inject) - else #inject (the (Symtab.lookup (dt_info, tname)))); + (if i < length newTs then List.nth (constr_inject, i) + else #inject (valOf (Symtab.lookup (dt_info, tname)))); fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) = let - val k = length (filter is_rec_type cargs) + val k = length (List.filter is_rec_type cargs) in (EVERY [DETERM tac, REPEAT (etac ex1E 1), rtac ex1I 1, @@ -189,7 +189,7 @@ intrs, j + 1) end; - val (tac', intrs', _) = foldl (mk_unique_constr_tac (length constrs)) + val (tac', intrs', _) = Library.foldl (mk_unique_constr_tac (length constrs)) ((tac, intrs, 0), constrs); in (tac', intrs') end; @@ -206,7 +206,7 @@ ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts); val induct' = cterm_instantiate ((map cert induct_Ps) ~~ (map cert insts)) induct; - val (tac, _) = foldl mk_unique_tac + val (tac, _) = Library.foldl mk_unique_tac (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_tac) 1 THEN rewtac (mk_meta_eq choice_eq), rec_intrs), descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts); @@ -255,7 +255,7 @@ in thy2 |> Theory.add_path (space_implode "_" new_type_names) |> PureThy.add_thmss [(("recs", rec_thms), [])] |>> - Theory.parent_path |> apsnd (pair reccomb_names o flat) + Theory.parent_path |> apsnd (pair reccomb_names o List.concat) end; @@ -267,10 +267,10 @@ val thy1 = add_path flat_names (space_implode "_" new_type_names) thy; - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val used = foldr add_typ_tfree_names (recTs, []); - val newTs = take (length (hd descr), recTs); + val used = Library.foldr add_typ_tfree_names (recTs, []); + val newTs = Library.take (length (hd descr), recTs); val T' = TFree (variant used "'t", HOLogic.typeS); fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T'; @@ -278,7 +278,7 @@ val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) => let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val Ts' = map mk_dummyT (filter is_rec_type cargs) + val Ts' = map mk_dummyT (List.filter is_rec_type cargs) in Const ("arbitrary", Ts @ Ts' ---> T') end) constrs) descr'; @@ -287,15 +287,15 @@ (* define case combinators via primrec combinators *) - val (case_defs, thy2) = foldl (fn ((defs, thy), + val (case_defs, thy2) = Library.foldl (fn ((defs, thy), ((((i, (_, _, constrs)), T), name), recname)) => let val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) => let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val Ts' = Ts @ map mk_dummyT (filter is_rec_type cargs); + val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs); val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts')); - val frees = take (length cargs, frees'); + val frees = Library.take (length cargs, frees'); val free = mk_Free "f" (Ts ---> T') j in (free, list_abs_free (map dest_Free frees', @@ -303,20 +303,20 @@ end) (constrs ~~ (1 upto length constrs))); val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T'; - val fns = (flat (take (i, case_dummy_fns))) @ - fns2 @ (flat (drop (i + 1, case_dummy_fns))); + val fns = (List.concat (Library.take (i, case_dummy_fns))) @ + fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))); val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T'); val decl = (Sign.base_name name, caseT, NoSyn); val def = ((Sign.base_name name) ^ "_def", Logic.mk_equals (list_comb (Const (name, caseT), fns1), - list_comb (reccomb, (flat (take (i, case_dummy_fns))) @ - fns2 @ (flat (drop (i + 1, case_dummy_fns))) ))); + list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @ + fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) ))); val (thy', [def_thm]) = thy |> Theory.add_consts_i [decl] |> (PureThy.add_defs_i false o map Thm.no_attributes) [def]; in (defs @ [def_thm], thy') end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~ - (take (length newTs, reccomb_names))); + (Library.take (length newTs, reccomb_names))); val case_thms = map (map (fn t => prove_goalw_cterm (case_defs @ (map mk_meta_eq primrec_thms)) (cterm_of (Theory.sign_of thy2) t) @@ -338,9 +338,9 @@ let val _ = message "Proving equations for case splitting ..."; - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val newTs = take (length (hd descr), recTs); + val newTs = Library.take (length (hd descr), recTs); fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) = @@ -372,7 +372,7 @@ fun prove_size_thms flat_names new_type_names descr sorts reccomb_names primrec_thms thy = if exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt => is_rec_type dt andalso not (null (fst (strip_dtyp dt)))) cargs) constrs) - (flat descr) + (List.concat descr) then (thy, []) else @@ -382,13 +382,13 @@ val big_name = space_implode "_" new_type_names; val thy1 = add_path flat_names big_name thy; - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; val size_name = "Nat.size"; val size_names = replicate (length (hd descr)) size_name @ map (Sign.full_name (Theory.sign_of thy1)) (DatatypeProp.indexify_names - (map (fn T => name_of_typ T ^ "_size") (drop (length (hd descr), recTs)))); + (map (fn T => name_of_typ T ^ "_size") (Library.drop (length (hd descr), recTs)))); val def_names = map (fn s => s ^ "_def") (DatatypeProp.indexify_names (map (fn T => name_of_typ T ^ "_size") recTs)); @@ -397,20 +397,20 @@ fun make_sizefun (_, cargs) = let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val k = length (filter is_rec_type cargs); + val k = length (List.filter is_rec_type cargs); val t = if k = 0 then HOLogic.zero else foldl1 plus (map Bound (k - 1 downto 0) @ [HOLogic.mk_nat 1]) in - foldr (fn (T, t') => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT, t) + Library.foldr (fn (T, t') => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT, t) end; - val fs = flat (map (fn (_, (_, _, constrs)) => map make_sizefun constrs) descr'); + val fs = List.concat (map (fn (_, (_, _, constrs)) => map make_sizefun constrs) descr'); val fTs = map fastype_of fs; val (thy', size_def_thms) = thy1 |> Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T --> HOLogic.natT, NoSyn)) - (drop (length (hd descr), size_names ~~ recTs))) |> + (Library.drop (length (hd descr), size_names ~~ recTs))) |> (PureThy.add_defs_i true o map Thm.no_attributes) (map (fn (((s, T), def_name), rec_name) => (def_name, Logic.mk_equals (Const (s, T --> HOLogic.natT), list_comb (Const (rec_name, fTs @ [T] ---> HOLogic.natT), fs)))) @@ -426,7 +426,7 @@ in thy' |> Theory.add_path big_name |> PureThy.add_thmss [(("size", size_thms), [])] |>> - Theory.parent_path |> apsnd flat + Theory.parent_path |> apsnd List.concat end; fun prove_weak_case_congs new_type_names descr sorts thy = diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_aux.ML --- a/src/HOL/Tools/datatype_aux.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_aux.ML Thu Mar 03 12:43:01 2005 +0100 @@ -71,7 +71,7 @@ fun message s = if !quiet_mode then () else writeln s; (* FIXME: move to library ? *) -fun foldl1 f (x::xs) = foldl f (x, xs); +fun foldl1 f (x::xs) = Library.foldl f (x, xs); fun add_path flat_names s = if flat_names then I else Theory.add_path s; fun parent_path flat_names = if flat_names then I else Theory.parent_path; @@ -108,7 +108,7 @@ fun cong_tac i st = (case Logic.strip_assums_concl - (nth_elem (i - 1, prems_of st)) of + (List.nth (prems_of st, i - 1)) of _ $ (_ $ (f $ x) $ (g $ y)) => let val cong' = lift_rule (st, i) cong; @@ -128,7 +128,7 @@ let val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule)); val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop - (Logic.strip_imp_concl (nth_elem (i - 1, prems_of st)))); + (Logic.strip_imp_concl (List.nth (prems_of st, i - 1)))); val getP = if can HOLogic.dest_imp (hd ts) then (apfst SOME) o HOLogic.dest_imp else pair NONE; fun abstr (t1, t2) = (case t1 of @@ -148,14 +148,14 @@ fun exh_tac exh_thm_of i state = let val sg = Thm.sign_of_thm state; - val prem = nth_elem (i - 1, prems_of state); + val prem = List.nth (prems_of state, i - 1); val params = Logic.strip_params prem; val (_, Type (tname, _)) = hd (rev params); val exhaustion = lift_rule (state, i) (exh_thm_of tname); val prem' = hd (prems_of exhaustion); val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem')); val exhaustion' = cterm_instantiate [(cterm_of sg (head_of lhs), - cterm_of sg (foldr (fn ((_, T), t) => Abs ("z", T, t)) + cterm_of sg (Library.foldr (fn ((_, T), t) => Abs ("z", T, t)) (params, Bound 0)))] exhaustion in compose_tac (false, exhaustion', nprems_of exhaustion) i state end; @@ -228,7 +228,7 @@ fun name_of_typ (Type (s, Ts)) = let val s' = Sign.base_name s - in space_implode "_" (filter (not o equal "") (map name_of_typ Ts) @ + in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @ [if Syntax.is_identifier s' then s' else "x"]) end | name_of_typ _ = ""; @@ -242,9 +242,9 @@ DtRec (find_index (curry op = tname o fst) new_dts) else error ("Illegal occurence of recursive type " ^ tname)); -fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, the (assoc (sorts, a))) +fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, valOf (assoc (sorts, a))) | typ_of_dtyp descr sorts (DtRec i) = - let val (s, ds, _) = the (assoc (descr, i)) + let val (s, ds, _) = valOf (assoc (descr, i)) in Type (s, map (typ_of_dtyp descr sorts) ds) end | typ_of_dtyp descr sorts (DtType (s, ds)) = Type (s, map (typ_of_dtyp descr sorts) ds); @@ -252,8 +252,8 @@ (* find all non-recursive types in datatype description *) fun get_nonrec_types descr sorts = - map (typ_of_dtyp descr sorts) (foldl (fn (Ts, (_, (_, _, constrs))) => - foldl (fn (Ts', (_, cargs)) => + map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) => + Library.foldl (fn (Ts', (_, cargs)) => filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr)); (* get all recursive types in datatype description *) @@ -264,22 +264,22 @@ (* get all branching types *) fun get_branching_types descr sorts = - map (typ_of_dtyp descr sorts) (foldl (fn (Ts, (_, (_, _, constrs))) => - foldl (fn (Ts', (_, cargs)) => foldr op union (map (fst o strip_dtyp) + map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) => + Library.foldl (fn (Ts', (_, cargs)) => Library.foldr op union (map (fst o strip_dtyp) cargs, Ts')) (Ts, constrs)) ([], descr)); -fun get_arities descr = foldl (fn (is, (_, (_, _, constrs))) => - foldl (fn (is', (_, cargs)) => map (length o fst o strip_dtyp) - (filter is_rec_type cargs) union is') (is, constrs)) ([], descr); +fun get_arities descr = Library.foldl (fn (is, (_, (_, _, constrs))) => + Library.foldl (fn (is', (_, cargs)) => map (length o fst o strip_dtyp) + (List.filter is_rec_type cargs) union is') (is, constrs)) ([], descr); (* nonemptiness check for datatypes *) fun check_nonempty descr = let - val descr' = flat descr; + val descr' = List.concat descr; fun is_nonempty_dt is i = let - val (_, _, constrs) = the (assoc (descr', i)); + val (_, _, constrs) = valOf (assoc (descr', i)); fun arg_nonempty (_, DtRec i) = if i mem is then false else is_nonempty_dt (i::is) i | arg_nonempty _ = true; @@ -303,8 +303,8 @@ NONE => typ_error T (tname ^ " is not a datatype - can't use it in\ \ nested recursion") | (SOME {index, descr, ...}) => - let val (_, vars, _) = the (assoc (descr, index)); - val subst = ((map dest_DtTFree vars) ~~ dts) handle LIST _ => + let val (_, vars, _) = valOf (assoc (descr, index)); + val subst = ((map dest_DtTFree vars) ~~ dts) handle UnequalLengths => typ_error T ("Type constructor " ^ tname ^ " used with wrong\ \ number of arguments") in (i + index, map (fn (j, (tn, args, cs)) => (i + j, @@ -333,18 +333,18 @@ (* unfold a constructor *) fun unfold_constr ((i, constrs, descrs), (cname, cargs)) = - let val (i', cargs', descrs') = foldl unfold_arg ((i, [], descrs), cargs) + let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs) in (i', constrs @ [(cname, cargs')], descrs') end; (* unfold a single datatype *) fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) = let val (i', constrs', descrs') = - foldl unfold_constr ((i, [], descrs), constrs) + Library.foldl unfold_constr ((i, [], descrs), constrs) in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs') end; - val (i', descr', descrs) = foldl unfold_datatype ((i, [],[]), descr); + val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr); in (descr' :: descrs, i') end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_codegen.ML --- a/src/HOL/Tools/datatype_codegen.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_codegen.ML Thu Mar 03 12:43:01 2005 +0100 @@ -17,7 +17,7 @@ fun mk_tuple [p] = p | mk_tuple ps = Pretty.block (Pretty.str "(" :: - flat (separate [Pretty.str ",", Pretty.brk 1] (map single ps)) @ + List.concat (separate [Pretty.str ",", Pretty.brk 1] (map single ps)) @ [Pretty.str ")"]); (**** datatype definition ****) @@ -26,16 +26,16 @@ fun find_nonempty (descr: DatatypeAux.descr) is i = let - val (_, _, constrs) = the (assoc (descr, i)); + val (_, _, constrs) = valOf (assoc (descr, i)); fun arg_nonempty (_, DatatypeAux.DtRec i) = if i mem is then NONE - else apsome (curry op + 1 o snd) (find_nonempty descr (i::is) i) + else Option.map (curry op + 1 o snd) (find_nonempty descr (i::is) i) | arg_nonempty _ = SOME 0; - fun max xs = foldl + fun max xs = Library.foldl (fn (NONE, _) => NONE | (SOME i, SOME j) => SOME (Int.max (i, j)) | (_, NONE) => NONE) (SOME 0, xs); val xs = sort (int_ord o pairself snd) - (mapfilter (fn (s, dts) => apsome (pair s) + (List.mapPartial (fn (s, dts) => Option.map (pair s) (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs) in case xs of [] => NONE | x :: _ => SOME x end; @@ -44,8 +44,8 @@ val sg = sign_of thy; val tab = DatatypePackage.get_datatypes thy; - val descr' = filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr; - val rtnames = map (#1 o snd) (filter (fn (_, (_, _, cs)) => + val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr; + val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) => exists (exists DatatypeAux.is_rec_type o snd) cs) descr'); val (_, (_, _, (cname, _) :: _)) :: _ = descr'; @@ -67,11 +67,11 @@ (if null tvs then [] else [mk_tuple (map Pretty.str tvs), Pretty.str " "]) @ [Pretty.str (mk_type_id sg tname ^ " ="), Pretty.brk 1] @ - flat (separate [Pretty.brk 1, Pretty.str "| "] + List.concat (separate [Pretty.brk 1, Pretty.str "| "] (map (fn (cname, ps') => [Pretty.block (Pretty.str (mk_const_id sg cname) :: (if null ps' then [] else - flat ([Pretty.str " of", Pretty.brk 1] :: + List.concat ([Pretty.str " of", Pretty.brk 1] :: separate [Pretty.str " *", Pretty.brk 1] (map single ps'))))]) ps))) :: rest) end; @@ -94,7 +94,7 @@ else parens (Pretty.block [Pretty.str (mk_const_id sg cname), Pretty.brk 1, mk_tuple args]), Pretty.str " =", Pretty.brk 1] @ - flat (separate [Pretty.str " $", Pretty.brk 1] + List.concat (separate [Pretty.str " $", Pretty.brk 1] ([Pretty.str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1, mk_type false (Ts ---> T), Pretty.str ")"] :: map (fn (x, U) => [Pretty.block [mk_term_of sg false U, @@ -108,7 +108,7 @@ val tvs = map DatatypeAux.dest_DtTFree dts; val sorts = map (rpair []) tvs; val (cs1, cs2) = - partition (exists DatatypeAux.is_rec_type o snd) cs; + List.partition (exists DatatypeAux.is_rec_type o snd) cs; val SOME (cname, _) = find_nonempty descr [i] i; fun mk_delay p = Pretty.block @@ -130,7 +130,7 @@ fun mk_choice [c] = mk_constr "(i-1)" false c | mk_choice cs = Pretty.block [Pretty.str "one_of", Pretty.brk 1, Pretty.blk (1, Pretty.str "[" :: - flat (separate [Pretty.str ",", Pretty.fbrk] + List.concat (separate [Pretty.str ",", Pretty.fbrk] (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @ [Pretty.str "]"]), Pretty.brk 1, Pretty.str "()"]; @@ -155,7 +155,7 @@ [Pretty.block [Pretty.str "(case", Pretty.brk 1, Pretty.str "i", Pretty.brk 1, Pretty.str "of", Pretty.brk 1, Pretty.str "0 =>", Pretty.brk 1, - mk_constr "0" true (cname, the (assoc (cs, cname))), + mk_constr "0" true (cname, valOf (assoc (cs, cname))), Pretty.brk 1, Pretty.str "| _ =>", Pretty.brk 1, mk_choice cs1, Pretty.str ")"]] else [mk_choice cs2])) :: @@ -199,18 +199,18 @@ invoke_codegen thy dep brack (gr, eta_expand c ts (i+1)) else let - val ts1 = take (i, ts); - val t :: ts2 = drop (i, ts); - val names = foldr add_term_names (ts1, - map (fst o fst o dest_Var) (foldr add_term_vars (ts1, []))); - val (Ts, dT) = split_last (take (i+1, fst (strip_type T))); + val ts1 = Library.take (i, ts); + val t :: ts2 = Library.drop (i, ts); + val names = Library.foldr add_term_names (ts1, + map (fst o fst o dest_Var) (Library.foldr add_term_vars (ts1, []))); + val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T))); fun pcase gr [] [] [] = ([], gr) | pcase gr ((cname, cargs)::cs) (t::ts) (U::Us) = let val j = length cargs; val xs = variantlist (replicate j "x", names); - val Us' = take (j, fst (strip_type U)); + val Us' = Library.take (j, fst (strip_type U)); val frees = map Free (xs ~~ Us'); val (gr0, cp) = invoke_codegen thy dep false (gr, list_comb (Const (cname, Us' ---> dT), frees)); @@ -222,7 +222,7 @@ end; val (ps1, gr1) = pcase gr constrs ts1 Ts; - val ps = flat (separate [Pretty.brk 1, Pretty.str "| "] ps1); + val ps = List.concat (separate [Pretty.brk 1, Pretty.str "| "] ps1); val (gr2, p) = invoke_codegen thy dep false (gr1, t); val (gr3, ps2) = foldl_map (invoke_codegen thy dep true) (gr2, ts2) in (gr3, (if not (null ts2) andalso brack then parens else I) @@ -257,15 +257,15 @@ (c as Const (s, T), ts) => (case find_first (fn (_, {index, descr, case_name, ...}) => s = case_name orelse - is_some (assoc (#3 (the (assoc (descr, index))), s))) + isSome (assoc (#3 (valOf (assoc (descr, index))), s))) (Symtab.dest (DatatypePackage.get_datatypes thy)) of NONE => NONE | SOME (tname, {index, descr, ...}) => - if is_some (get_assoc_code thy s T) then NONE else + if isSome (get_assoc_code thy s T) then NONE else let val SOME (_, _, constrs) = assoc (descr, index) in (case (assoc (constrs, s), strip_type T) of (NONE, _) => SOME (pretty_case thy gr dep brack - (#3 (the (assoc (descr, index)))) c ts) + (#3 (valOf (assoc (descr, index)))) c ts) | (SOME args, (_, Type _)) => SOME (pretty_constr thy (fst (invoke_tycodegen thy dep false (gr, snd (strip_type T)))) dep brack args c ts) @@ -277,7 +277,7 @@ (case Symtab.lookup (DatatypePackage.get_datatypes thy, s) of NONE => NONE | SOME {descr, ...} => - if is_some (get_assoc_type thy s) then NONE else + if isSome (get_assoc_type thy s) then NONE else let val (gr', ps) = foldl_map (invoke_tycodegen thy dep false) (gr, Ts) in SOME (add_dt_defs thy dep gr' descr, diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_package.ML --- a/src/HOL/Tools/datatype_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -125,15 +125,15 @@ fun constrs_of_sg sg tname = (case datatype_info_sg sg tname of SOME {index, descr, ...} => - let val (_, _, constrs) = the (assoc (descr, index)) - in SOME (map (fn (cname, _) => Const (cname, the (Sign.const_type sg cname))) constrs) + let val (_, _, constrs) = valOf (assoc (descr, index)) + in SOME (map (fn (cname, _) => Const (cname, valOf (Sign.const_type sg cname))) constrs) end | _ => NONE); val constrs_of = constrs_of_sg o Theory.sign_of; fun case_const_of thy tname = (case datatype_info thy tname of - SOME {case_name, ...} => SOME (Const (case_name, the (Sign.const_type + SOME {case_name, ...} => SOME (Const (case_name, valOf (Sign.const_type (Theory.sign_of thy) case_name))) | _ => NONE); @@ -169,7 +169,7 @@ fun occs_in_prems tacf vars = SUBGOAL (fn (Bi, i) => (if exists (fn Free (a, _) => a mem vars) - (foldr add_term_frees (#2 (strip_context Bi), [])) + (Library.foldr add_term_frees (#2 (strip_context Bi), [])) then warning "Induction variable occurs also among premises!" else (); tacf i)); @@ -182,9 +182,9 @@ fun prep_var (Var (ixn, _), SOME x) = SOME (ixn, x) | prep_var _ = NONE; -fun prep_inst (concl, xs) = (*exception LIST*) +fun prep_inst (concl, xs) = (*exception UnequalLengths *) let val vs = InductAttrib.vars_of concl - in mapfilter prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end; + in List.mapPartial prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end; in @@ -196,11 +196,11 @@ (case opt_rule of SOME r => (r, "Induction rule") | NONE => - let val tn = find_tname (hd (mapfilter I (flat varss))) Bi + let val tn = find_tname (hd (List.mapPartial I (List.concat varss))) Bi in (#induction (datatype_info_sg_err sign tn), "Induction rule for type " ^ tn) end); val concls = HOLogic.dest_concls (Thm.concl_of rule); - val insts = flat (map prep_inst (concls ~~ varss)) handle LIST _ => + val insts = List.concat (map prep_inst (concls ~~ varss)) handle UnequalLengths => error (rule_name ^ " has different numbers of variables"); in occs_in_prems (inst_tac insts rule) (map #2 insts) i state end; @@ -272,20 +272,20 @@ local fun dt_recs (DtTFree _) = [] - | dt_recs (DtType (_, dts)) = flat (map dt_recs dts) + | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts) | dt_recs (DtRec i) = [i]; fun dt_cases (descr: descr) (_, args, constrs) = let - fun the_bname i = Sign.base_name (#1 (the (assoc (descr, i)))); - val bnames = map the_bname (distinct (flat (map dt_recs args))); + fun the_bname i = Sign.base_name (#1 (valOf (assoc (descr, i)))); + val bnames = map the_bname (distinct (List.concat (map dt_recs args))); in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end; fun induct_cases descr = - DatatypeProp.indexify_names (flat (map (dt_cases descr) (map #2 descr))); + DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr))); -fun exhaust_cases descr i = dt_cases descr (the (assoc (descr, i))); +fun exhaust_cases descr i = dt_cases descr (valOf (assoc (descr, i))); in @@ -293,18 +293,18 @@ fun mk_case_names_exhausts descr new = map (RuleCases.case_names o exhaust_cases descr o #1) - (filter (fn ((_, (name, _, _))) => name mem_string new) descr); + (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); end; fun add_rules simps case_thms size_thms rec_thms inject distinct weak_case_congs cong_att = (#1 o PureThy.add_thmss [(("simps", simps), []), - (("", flat case_thms @ size_thms @ - flat distinct @ rec_thms), [Simplifier.simp_add_global]), + (("", List.concat case_thms @ size_thms @ + List.concat distinct @ rec_thms), [Simplifier.simp_add_global]), (("", size_thms @ rec_thms), [RecfunCodegen.add]), - (("", flat inject), [iff_add_global]), - (("", flat distinct RL [notE]), [Classical.safe_elim_global]), + (("", List.concat inject), [iff_add_global]), + (("", List.concat distinct RL [notE]), [Classical.safe_elim_global]), (("", weak_case_congs), [cong_att])]); @@ -325,7 +325,7 @@ (("", exhaustion), [InductAttrib.cases_type_global name])]; fun unnamed_rule i = (("", proj i induction), [InductAttrib.induct_type_global ""]); - val rules = flat (map named_rules infos) @ map unnamed_rule (length infos upto n - 1); + val rules = List.concat (map named_rules infos) @ map unnamed_rule (length infos upto n - 1); in #1 o PureThy.add_thms rules end; @@ -393,7 +393,7 @@ fun case_tr sg [t, u] = let fun case_error s name ts = raise TERM ("Error in case expression" ^ - if_none (apsome (curry op ^ " for datatype ") name) "" ^ ":\n" ^ s, ts); + getOpt (Option.map (curry op ^ " for datatype ") name, "") ^ ":\n" ^ s, ts); fun dest_case1 (Const ("_case1", _) $ t $ u) = (case strip_comb t of (Const (s, _), ts) => (Sign.intern_const sg s, ts) | (Free (s, _), ts) => (Sign.intern_const sg s, ts) @@ -413,13 +413,13 @@ Syntax.const "split" $ abstr (x, abstr (y, body)) | abstr (t, _) = case_error "Illegal pattern" NONE [t]; in case find_first (fn (_, {descr, index, ...}) => - exists (equal cname o fst) (#3 (snd (nth_elem (index, descr))))) tab of + exists (equal cname o fst) (#3 (snd (List.nth (descr, index))))) tab of NONE => case_error ("Not a datatype constructor: " ^ cname) NONE [u] | SOME (tname, {descr, case_name, index, ...}) => let val _ = if exists (equal "dummy_pattern" o fst o fst) cases' then case_error "Illegal occurrence of '_' dummy pattern" (SOME tname) [u] else (); - val (_, (_, dts, constrs)) = nth_elem (index, descr); + val (_, (_, dts, constrs)) = List.nth (descr, index); val sorts = map (rpair [] o dest_DtTFree) dts; fun find_case (cases, (s, dt)) = (case find_first (equal s o fst o fst) cases' of @@ -431,7 +431,7 @@ if length dt <> length vs then case_error ("Wrong number of arguments for constructor " ^ s) (SOME tname) vs - else (cases \ c, foldr abstr (vs, t))) + else (cases \ c, Library.foldr abstr (vs, t))) val (cases'', fs) = foldl_map find_case (cases', constrs) in case (cases'', length constrs = length cases', default) of ([], true, SOME _) => @@ -473,7 +473,7 @@ NONE => (body, [cname]) :: cs | SOME cnames => overwrite (cs, (body, cnames @ [cname]))); val cases' = sort (int_ord o Library.swap o pairself (length o snd)) - (foldl count_cases ([], cases)); + (Library.foldl count_cases ([], cases)); fun mk_case1 (cname, (vs, body), _) = Syntax.const "_case1" $ list_comb (Syntax.const cname, vs) $ body; in @@ -490,7 +490,7 @@ [("dummy_pattern", ([], default), false)])) end; -fun make_case_tr' case_names descr = flat (map +fun make_case_tr' case_names descr = List.concat (map (fn ((_, (_, _, constrs)), case_name) => map (rpair (case_tr' constrs)) (NameSpace.accesses' case_name)) (descr ~~ case_names)); @@ -542,7 +542,7 @@ (********************* axiomatic introduction of datatypes ********************) fun add_and_get_axioms_atts label tnames attss ts thy = - foldr (fn (((tname, atts), t), (thy', axs)) => + Library.foldr (fn (((tname, atts), t), (thy', axs)) => let val (thy'', [ax]) = thy' |> Theory.add_path tname |> @@ -554,7 +554,7 @@ add_and_get_axioms_atts label tnames (replicate (length tnames) []); fun add_and_get_axiomss label tnames tss thy = - foldr (fn ((tname, ts), (thy', axss)) => + Library.foldr (fn ((tname, ts), (thy', axss)) => let val (thy'', [axs]) = thy' |> Theory.add_path tname |> @@ -565,10 +565,10 @@ fun add_datatype_axm flat_names new_type_names descr sorts types_syntax constr_syntax dt_info case_names_induct case_names_exhausts thy = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val used = foldr add_typ_tfree_names (recTs, []); - val newTs = take (length (hd descr), recTs); + val used = Library.foldr add_typ_tfree_names (recTs, []); + val newTs = Library.take (length (hd descr), recTs); val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt => is_rec_type dt andalso not (null (fst (strip_dtyp dt)))) @@ -591,7 +591,7 @@ (1 upto (length descr'))); val size_names = DatatypeProp.indexify_names - (map (fn T => name_of_typ T ^ "_size") (drop (length (hd descr), recTs))); + (map (fn T => name_of_typ T ^ "_size") (Library.drop (length (hd descr), recTs))); val freeT = TFree (variant used "'t", HOLogic.typeS); val case_fn_Ts = map (fn (i, (_, _, constrs)) => @@ -605,7 +605,7 @@ (** new types **) - curry (foldr (fn (((name, mx), tvs), thy') => thy' |> + curry (Library.foldr (fn (((name, mx), tvs), thy') => thy' |> TypedefPackage.add_typedecls [(name, tvs, mx)])) (types_syntax ~~ tyvars) |> add_path flat_names (space_implode "_" new_type_names) |> @@ -631,12 +631,12 @@ (if no_size then I else Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T --> HOLogic.natT, NoSyn)) - (size_names ~~ drop (length (hd descr), recTs)))) |> + (size_names ~~ Library.drop (length (hd descr), recTs)))) |> (** constructors **) parent_path flat_names |> - curry (foldr (fn (((((_, (_, _, constrs)), T), tname), + curry (Library.foldr (fn (((((_, (_, _, constrs)), T), tname), constr_syntax'), thy') => thy' |> add_path flat_names tname |> Theory.add_consts_i (map (fn ((_, cargs), (cname, mx)) => @@ -686,7 +686,7 @@ exhaustion ~~ replicate (length (hd descr)) QuickAndDirty ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs); - val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms; + val simps = List.concat (distinct @ inject @ case_thms) @ size_thms @ rec_thms; val split_thms = split ~~ split_asm; val thy12 = thy11 |> @@ -694,7 +694,7 @@ Theory.add_path (space_implode "_" new_type_names) |> add_rules simps case_thms size_thms rec_thms inject distinct weak_case_congs Simplifier.cong_add_global |> - put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |> + put_datatypes (Library.foldr Symtab.update (dt_infos, dt_info)) |> add_cases_induct dt_infos induct |> Theory.parent_path |> (#1 o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)) |> @@ -741,18 +741,18 @@ val (thy11, size_thms) = DatatypeAbsProofs.prove_size_thms flat_names new_type_names descr sorts reccomb_names rec_thms thy10; - val dt_infos = map (make_dt_info (flat descr) induct reccomb_names rec_thms) + val dt_infos = map (make_dt_info (List.concat descr) induct reccomb_names rec_thms) ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~ casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs); - val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms; + val simps = List.concat (distinct @ inject @ case_thms) @ size_thms @ rec_thms; val thy12 = thy11 |> Theory.add_advanced_trfuns ([], [], make_case_tr' case_names (hd descr), []) |> Theory.add_path (space_implode "_" new_type_names) |> add_rules simps case_thms size_thms rec_thms inject distinct weak_case_congs (Simplifier.change_global_ss (op addcongs)) |> - put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |> + put_datatypes (Library.foldr Symtab.update (dt_infos, dt_info)) |> add_cases_induct dt_infos induct |> Theory.parent_path |> (#1 o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)) |> @@ -778,7 +778,7 @@ val _ = Theory.requires thy0 "Inductive" "datatype representations"; fun app_thmss srcs thy = foldl_map (fn (thy, x) => apply_theorems x thy) (thy, srcs); - fun app_thm src thy = apsnd Library.hd (apply_theorems [src] thy); + fun app_thm src thy = apsnd hd (apply_theorems [src] thy); val (((thy1, induction), inject), distinct) = thy0 |> app_thmss raw_distinct @@ -796,7 +796,7 @@ | get_typ t = err t; val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induction'))); - val new_type_names = if_none alt_names (map fst dtnames); + val new_type_names = getOpt (alt_names, map fst dtnames); fun get_constr t = (case Logic.strip_assums_concl t of _ $ (_ $ t') => (case head_of t' of @@ -854,13 +854,13 @@ ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~ map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs); - val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms; + val simps = List.concat (distinct @ inject @ case_thms) @ size_thms @ rec_thms; val thy11 = thy10 |> Theory.add_advanced_trfuns ([], [], make_case_tr' case_names descr, []) |> add_rules simps case_thms size_thms rec_thms inject distinct weak_case_congs (Simplifier.change_global_ss (op addcongs)) |> - put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |> + put_datatypes (Library.foldr Symtab.update (dt_infos, dt_info)) |> add_cases_induct dt_infos induction' |> Theory.parent_path |> (#1 o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)) |> @@ -915,8 +915,8 @@ let fun prep_constr ((constrs, constr_syntax', sorts'), (cname, cargs, mx')) = let - val (cargs', sorts'') = foldl (prep_typ sign) (([], sorts'), cargs); - val _ = (case foldr add_typ_tfree_names (cargs', []) \\ tvs of + val (cargs', sorts'') = Library.foldl (prep_typ sign) (([], sorts'), cargs); + val _ = (case Library.foldr add_typ_tfree_names (cargs', []) \\ tvs of [] => () | vs => error ("Extra type variables on rhs: " ^ commas vs)) in (constrs @ [((if flat_names then Sign.full_name sign else @@ -928,7 +928,7 @@ " of datatype " ^ tname); val (constrs', constr_syntax', sorts') = - foldl prep_constr (([], [], sorts), constrs) + Library.foldl prep_constr (([], [], sorts), constrs) in case duplicates (map fst constrs') of @@ -940,7 +940,7 @@ " in datatype " ^ tname) end; - val (dts', constr_syntax, sorts', i) = foldl prep_dt_spec (([], [], [], 0), dts); + val (dts', constr_syntax, sorts', i) = Library.foldl prep_dt_spec (([], [], [], 0), dts); val sorts = sorts' @ (map (rpair (Sign.defaultS sign)) (tyvars \\ map fst sorts')); val dt_info = get_datatypes thy; val (descr, _) = unfold_datatypes sign dts' sorts dt_info dts' i; @@ -948,7 +948,7 @@ if err then error ("NONEmptiness check failed for datatype " ^ s) else raise exn; - val descr' = flat descr; + val descr' = List.concat descr; val case_names_induct = mk_case_names_induct descr'; val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts); in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_prop.ML --- a/src/HOL/Tools/datatype_prop.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_prop.ML Thu Mar 03 12:43:01 2005 +0100 @@ -64,7 +64,7 @@ fun make_injs descr sorts = let - val descr' = flat descr; + val descr' = List.concat descr; fun make_inj T ((cname, cargs), injs) = if null cargs then injs else @@ -80,22 +80,22 @@ (map HOLogic.mk_eq (frees ~~ frees')))))::injs end; - in map (fn (d, T) => foldr (make_inj T) (#3 (snd d), [])) - ((hd descr) ~~ take (length (hd descr), get_rec_types descr' sorts)) + in map (fn (d, T) => Library.foldr (make_inj T) (#3 (snd d), [])) + ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts)) end; (********************************* induction **********************************) fun make_ind descr sorts = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; val pnames = if length descr' = 1 then ["P"] else map (fn i => "P" ^ string_of_int i) (1 upto length descr'); fun make_pred i T = let val T' = T --> HOLogic.boolT - in Free (nth_elem (i, pnames), T') end; + in Free (List.nth (pnames, i), T') end; fun make_ind_prem k T (cname, cargs) = let @@ -105,11 +105,11 @@ (make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us))) end; - val recs = filter is_rec_type cargs; + val recs = List.filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr' sorts) cargs; val recTs' = map (typ_of_dtyp descr' sorts) recs; val tnames = variantlist (make_tnames Ts, pnames); - val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); + val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = tnames ~~ Ts; val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); @@ -118,7 +118,7 @@ list_comb (Const (cname, Ts ---> T), map Free frees)))) end; - val prems = flat (map (fn ((i, (_, _, constrs)), T) => + val prems = List.concat (map (fn ((i, (_, _, constrs)), T) => map (make_ind_prem i T) constrs) (descr' ~~ recTs)); val tnames = make_tnames recTs; val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") @@ -131,7 +131,7 @@ fun make_casedists descr sorts = let - val descr' = flat descr; + val descr' = List.concat descr; fun make_casedist_prem T (cname, cargs) = let @@ -149,29 +149,29 @@ end in map make_casedist - ((hd descr) ~~ take (length (hd descr), get_rec_types descr' sorts)) + ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts)) end; (*************** characteristic equations for primrec combinator **************) fun make_primrec_Ts descr sorts used = let - val descr' = flat descr; + val descr' = List.concat descr; val rec_result_Ts = map TFree (variantlist (replicate (length descr') "'t", used) ~~ replicate (length descr') HOLogic.typeS); - val reccomb_fn_Ts = flat (map (fn (i, (_, _, constrs)) => + val reccomb_fn_Ts = List.concat (map (fn (i, (_, _, constrs)) => map (fn (_, cargs) => let val Ts = map (typ_of_dtyp descr' sorts) cargs; - val recs = filter (is_rec_type o fst) (cargs ~~ Ts); + val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts); fun mk_argT (dt, T) = - binder_types T ---> nth_elem (body_index dt, rec_result_Ts); + binder_types T ---> List.nth (rec_result_Ts, body_index dt); val argTs = Ts @ map mk_argT recs - in argTs ---> nth_elem (i, rec_result_Ts) + in argTs ---> List.nth (rec_result_Ts, i) end) constrs) descr'); in (rec_result_Ts, reccomb_fn_Ts) end; @@ -180,9 +180,9 @@ let val sign = Theory.sign_of thy; - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val used = foldr add_typ_tfree_names (recTs, []); + val used = Library.foldr add_typ_tfree_names (recTs, []); val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used; @@ -200,18 +200,18 @@ fun make_primrec T comb_t ((ts, f::fs), (cname, cargs)) = let - val recs = filter is_rec_type cargs; + val recs = List.filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr' sorts) cargs; val recTs' = map (typ_of_dtyp descr' sorts) recs; val tnames = make_tnames Ts; - val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); + val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = map Free (tnames ~~ Ts); val frees' = map Free (rec_tnames ~~ recTs'); fun mk_reccomb ((dt, T), t) = let val (Us, U) = strip_type T in list_abs (map (pair "x") Us, - nth_elem (body_index dt, reccombs) $ app_bnds t (length Us)) + List.nth (reccombs, body_index dt) $ app_bnds t (length Us)) end; val reccombs' = map mk_reccomb (recs ~~ recTs' ~~ frees') @@ -221,8 +221,8 @@ list_comb (f, frees @ reccombs')))], fs) end - in fst (foldl (fn (x, ((dt, T), comb_t)) => - foldl (make_primrec T comb_t) (x, #3 (snd dt))) + in fst (Library.foldl (fn (x, ((dt, T), comb_t)) => + Library.foldl (make_primrec T comb_t) (x, #3 (snd dt))) (([], rec_fns), descr' ~~ recTs ~~ reccombs)) end; @@ -230,10 +230,10 @@ fun make_case_combs new_type_names descr sorts thy fname = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val used = foldr add_typ_tfree_names (recTs, []); - val newTs = take (length (hd descr), recTs); + val used = Library.foldr add_typ_tfree_names (recTs, []); + val newTs = Library.take (length (hd descr), recTs); val T' = TFree (variant used "'t", HOLogic.typeS); val case_fn_Ts = map (fn (i, (_, _, constrs)) => @@ -254,9 +254,9 @@ fun make_cases new_type_names descr sorts thy = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val newTs = take (length (hd descr), recTs); + val newTs = Library.take (length (hd descr), recTs); fun make_case T comb_t ((cname, cargs), f) = let @@ -276,9 +276,9 @@ fun make_distincts new_type_names descr sorts thy = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val newTs = take (length (hd descr), recTs); + val newTs = Library.take (length (hd descr), recTs); (**** number of constructors < dtK : C_i ... ~= C_j ... ****) @@ -315,10 +315,10 @@ fun make_splits new_type_names descr sorts thy = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val used' = foldr add_typ_tfree_names (recTs, []); - val newTs = take (length (hd descr), recTs); + val used' = Library.foldr add_typ_tfree_names (recTs, []); + val newTs = Library.take (length (hd descr), recTs); val T' = TFree (variant used' "'t", HOLogic.typeS); val P = Free ("P", T' --> HOLogic.boolT); @@ -334,13 +334,13 @@ val eqn = HOLogic.mk_eq (Free ("x", T), list_comb (Const (cname, Ts ---> T), frees)); val P' = P $ list_comb (f, frees) - in ((foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t)) + in ((Library.foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t)) (frees, HOLogic.imp $ eqn $ P'))::t1s, - (foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t)) + (Library.foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t)) (frees, HOLogic.conj $ eqn $ (HOLogic.Not $ P')))::t2s) end; - val (t1s, t2s) = foldr process_constr (constrs ~~ fs, ([], [])); + val (t1s, t2s) = Library.foldr process_constr (constrs ~~ fs, ([], [])); val lhs = P $ (comb_t $ Free ("x", T)) in (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, mk_conj t1s)), @@ -356,13 +356,13 @@ fun make_size descr sorts thy = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; val size_name = "Nat.size"; val size_names = replicate (length (hd descr)) size_name @ map (Sign.intern_const (Theory.sign_of thy)) (indexify_names - (map (fn T => name_of_typ T ^ "_size") (drop (length (hd descr), recTs)))); + (map (fn T => name_of_typ T ^ "_size") (Library.drop (length (hd descr), recTs)))); val size_consts = map (fn (s, T) => Const (s, T --> HOLogic.natT)) (size_names ~~ recTs); @@ -370,12 +370,12 @@ fun make_size_eqn size_const T (cname, cargs) = let - val recs = filter is_rec_type cargs; + val recs = List.filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr' sorts) cargs; val recTs = map (typ_of_dtyp descr' sorts) recs; val tnames = make_tnames Ts; - val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); - val ts = map (fn ((r, s), T) => nth_elem (dest_DtRec r, size_consts) $ + val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); + val ts = map (fn ((r, s), T) => List.nth (size_consts, dest_DtRec r) $ Free (s, T)) (recs ~~ rec_tnames ~~ recTs); val t = if ts = [] then HOLogic.zero else foldl1 plus (ts @ [HOLogic.mk_nat 1]) @@ -385,7 +385,7 @@ end in - flat (map (fn (((_, (_, _, constrs)), size_const), T) => + List.concat (map (fn (((_, (_, _, constrs)), size_const), T) => map (make_size_eqn size_const T) constrs) (descr' ~~ size_consts ~~ recTs)) end; @@ -465,9 +465,9 @@ fun make_nchotomys descr sorts = let - val descr' = flat descr; + val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; - val newTs = take (length (hd descr), recTs); + val newTs = Library.take (length (hd descr), recTs); fun mk_eqn T (cname, cargs) = let @@ -475,7 +475,7 @@ val tnames = variantlist (make_tnames Ts, ["v"]); val frees = tnames ~~ Ts in - foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t)) + Library.foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t)) (frees, HOLogic.mk_eq (Free ("v", T), list_comb (Const (cname, Ts ---> T), map Free frees))) end diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_realizer.ML --- a/src/HOL/Tools/datatype_realizer.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_realizer.ML Thu Mar 03 12:43:01 2005 +0100 @@ -53,23 +53,23 @@ fun make_pred i T U r x = if i mem is then - Free (nth_elem (i, pnames), T --> U --> HOLogic.boolT) $ r $ x - else Free (nth_elem (i, pnames), U --> HOLogic.boolT) $ x; + Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x + else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x; fun mk_all i s T t = if i mem is then list_all_free ([(s, T)], t) else t; - val (prems, rec_fns) = split_list (flat (snd (foldl_map + val (prems, rec_fns) = split_list (List.concat (snd (foldl_map (fn (j, ((i, (_, _, constrs)), T)) => foldl_map (fn (j, (cname, cargs)) => let val Ts = map (typ_of_dtyp descr sorts) cargs; val tnames = variantlist (DatatypeProp.make_tnames Ts, pnames); - val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts); + val recs = List.filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts); val frees = tnames ~~ Ts; fun mk_prems vs [] = let - val rT = nth_elem (i, rec_result_Ts); + val rT = List.nth (rec_result_Ts, i); val vs' = filter_out is_unit vs; val f = mk_Free "f" (map fastype_of vs' ---> rT) j; val f' = Pattern.eta_contract (list_abs_free @@ -83,7 +83,7 @@ val k = body_index dt; val (Us, U) = strip_type T; val i = length Us; - val rT = nth_elem (k, rec_result_Ts); + val rT = List.nth (rec_result_Ts, k); val r = Free ("r" ^ s, Us ---> rT); val (p, f) = mk_prems (vs @ [r]) ds in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies @@ -107,7 +107,7 @@ (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0))) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names); val r = if null is then Extraction.nullt else - foldr1 HOLogic.mk_prod (mapfilter (fn (((((i, _), T), U), s), tname) => + foldr1 HOLogic.mk_prod (List.mapPartial (fn (((((i, _), T), U), s), tname) => if i mem is then SOME (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T)) else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames)); @@ -130,22 +130,22 @@ val {path, ...} = Sign.rep_sg sg; val ind_name = Thm.name_of_thm induction; - val vs = map (fn i => nth_elem (i, pnames)) is; + val vs = map (fn i => List.nth (pnames, i)) is; val (thy', thm') = thy |> Theory.absolute_path |> PureThy.store_thm ((space_implode "_" (ind_name :: vs @ ["correctness"]), thm), []) - |>> Theory.add_path (NameSpace.pack (if_none path [])); + |>> Theory.add_path (NameSpace.pack (getOpt (path,[]))); val ivs = Drule.vars_of_terms [Logic.varify (DatatypeProp.make_ind [descr] sorts)]; val rvs = Drule.vars_of_terms [prop_of thm']; val ivs1 = map Var (filter_out (fn (_, T) => tname_of (body_type T) mem ["set", "bool"]) ivs); - val ivs2 = map (fn (ixn, _) => Var (ixn, the (assoc (rvs, ixn)))) ivs; + val ivs2 = map (fn (ixn, _) => Var (ixn, valOf (assoc (rvs, ixn)))) ivs; - val prf = foldr forall_intr_prf (ivs2, - foldr (fn ((f, p), prf) => + val prf = Library.foldr forall_intr_prf (ivs2, + Library.foldr (fn ((f, p), prf) => (case head_of (strip_abs_body f) of Free (s, T) => let val T' = Type.varifyT T @@ -156,7 +156,7 @@ (rec_fns ~~ prems_of thm, Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0)))); - val r' = if null is then r else Logic.varify (foldr (uncurry lambda) + val r' = if null is then r else Logic.varify (Library.foldr (uncurry lambda) (map Logic.unvarify ivs1 @ filter_out is_unit (map (head_of o strip_abs_body) rec_fns), r)); @@ -184,7 +184,7 @@ end; val SOME (_, _, constrs) = assoc (descr, index); - val T = nth_elem (index, get_rec_types descr sorts); + val T = List.nth (get_rec_types descr sorts, index); val (rs, prems) = split_list (map (make_casedist_prem T) constrs); val r = Const (case_name, map fastype_of rs ---> T --> rT); @@ -205,11 +205,11 @@ val (thy', thm') = thy |> Theory.absolute_path |> PureThy.store_thm ((exh_name ^ "_P_correctness", thm), []) - |>> Theory.add_path (NameSpace.pack (if_none path [])); + |>> Theory.add_path (NameSpace.pack (getOpt (path,[]))); val P = Var (("P", 0), rT' --> HOLogic.boolT); val prf = forall_intr_prf (y, forall_intr_prf (P, - foldr (fn ((p, r), prf) => + Library.foldr (fn ((p, r), prf) => forall_intr_prf (Logic.varify r, AbsP ("H", SOME (Logic.varify p), prf))) (prems ~~ rs, Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0))))); @@ -225,8 +225,8 @@ fun add_dt_realizers sorts infos thy = if !proofs < 2 then thy else (message "Adding realizers for induction and case analysis ..."; thy - |> curry (foldr (make_ind sorts (hd infos))) + |> curry (Library.foldr (make_ind sorts (hd infos))) (subsets 0 (length (#descr (hd infos)) - 1)) - |> curry (foldr (make_casedists sorts)) infos); + |> curry (Library.foldr (make_casedists sorts)) infos); end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/datatype_rep_proofs.ML --- a/src/HOL/Tools/datatype_rep_proofs.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/datatype_rep_proofs.ML Thu Mar 03 12:43:01 2005 +0100 @@ -35,7 +35,7 @@ fun exh_thm_of (dt_info : datatype_info Symtab.table) tname = - #exhaustion (the (Symtab.lookup (dt_info, tname))); + #exhaustion (valOf (Symtab.lookup (dt_info, tname))); (******************************************************************************) @@ -62,7 +62,7 @@ "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0", "Lim_inject", "Suml_inject", "Sumr_inject"]; - val descr' = flat descr; + val descr' = List.concat descr; val big_name = space_implode "_" new_type_names; val thy1 = add_path flat_names big_name thy; @@ -78,11 +78,11 @@ val branchT = if null branchTs then HOLogic.unitT else fold_bal (fn (T, U) => Type ("+", [T, U])) branchTs; val arities = get_arities descr' \ 0; - val unneeded_vars = hd tyvars \\ foldr add_typ_tfree_names (leafTs' @ branchTs, []); - val leafTs = leafTs' @ (map (fn n => TFree (n, the (assoc (sorts, n)))) unneeded_vars); + val unneeded_vars = hd tyvars \\ Library.foldr add_typ_tfree_names (leafTs' @ branchTs, []); + val leafTs = leafTs' @ (map (fn n => TFree (n, valOf (assoc (sorts, n)))) unneeded_vars); val recTs = get_rec_types descr' sorts; - val newTs = take (length (hd descr), recTs); - val oldTs = drop (length (hd descr), recTs); + val newTs = Library.take (length (hd descr), recTs); + val oldTs = Library.drop (length (hd descr), recTs); val sumT = if null leafTs then HOLogic.unitT else fold_bal (fn (T, U) => Type ("+", [T, U])) leafTs; val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT])); @@ -134,7 +134,7 @@ in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs) end; - val mk_lim = foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t))); + val mk_lim = Library.foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t))); (************** generate introduction rules for representing set **********) @@ -152,7 +152,7 @@ app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts) in (j + 1, list_all (map (pair "x") Ts, HOLogic.mk_Trueprop (HOLogic.mk_mem (free_t, - Const (nth_elem (k, rep_set_names), UnivT)))) :: prems, + Const (List.nth (rep_set_names, k), UnivT)))) :: prems, mk_lim (Ts, free_t) :: ts) end | _ => @@ -160,7 +160,7 @@ in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts) end); - val (_, prems, ts) = foldr mk_prem (cargs, (1, [], [])); + val (_, prems, ts) = Library.foldr mk_prem (cargs, (1, [], [])); val concl = HOLogic.mk_Trueprop (HOLogic.mk_mem (mk_univ_inj ts n i, Const (s, UnivT))) in Logic.list_implies (prems, concl) @@ -168,7 +168,7 @@ val consts = map (fn s => Const (s, UnivT)) rep_set_names; - val intr_ts = flat (map (fn ((_, (_, _, constrs)), rep_set_name) => + val intr_ts = List.concat (map (fn ((_, (_, _, constrs)), rep_set_name) => map (make_intr rep_set_name (length constrs)) ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names)); @@ -179,21 +179,21 @@ (********************************* typedef ********************************) - val thy3 = add_path flat_names big_name (foldl (fn (thy, ((((name, mx), tvs), c), name')) => + val thy3 = add_path flat_names big_name (Library.foldl (fn (thy, ((((name, mx), tvs), c), name')) => setmp TypedefPackage.quiet_mode true (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) c NONE (rtac exI 1 THEN QUIET_BREADTH_FIRST (has_fewer_prems 1) (resolve_tac rep_intrs 1))) thy |> #1) (parent_path flat_names thy2, types_syntax ~~ tyvars ~~ - (take (length newTs, consts)) ~~ new_type_names)); + (Library.take (length newTs, consts)) ~~ new_type_names)); (*********************** definition of constructors ***********************) val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_"; val rep_names = map (curry op ^ "Rep_") new_type_names; val rep_names' = map (fn i => big_rep_name ^ (string_of_int i)) - (1 upto (length (flat (tl descr)))); + (1 upto (length (List.concat (tl descr)))); val all_rep_names = map (Sign.intern_const (Theory.sign_of thy3)) rep_names @ map (Sign.full_name (Theory.sign_of thy3)) rep_names'; @@ -211,12 +211,12 @@ val free_t = mk_Free "x" T j in (case (strip_dtyp dt, strip_type T) of ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim (Us, - Const (nth_elem (m, all_rep_names), U --> Univ_elT) $ + Const (List.nth (all_rep_names, m), U --> Univ_elT) $ app_bnds free_t (length Us)) :: r_args) | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args)) end; - val (_, l_args, r_args) = foldr constr_arg (cargs, (1, [], [])); + val (_, l_args, r_args) = Library.foldr constr_arg (cargs, (1, [], [])); val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T; val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname); val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname); @@ -243,14 +243,14 @@ (Const (Sign.intern_const sg ("Rep_" ^ tname), T --> Univ_elT)); val cong' = standard (cterm_instantiate [(cterm_of sg cong_f, rep_const)] arg_cong); val dist = standard (cterm_instantiate [(cterm_of sg distinct_f, rep_const)] distinct_lemma); - val (thy', defs', eqns', _) = foldl ((make_constr_def tname T) (length constrs)) + val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs)) ((add_path flat_names tname thy, defs, [], 1), constrs ~~ constr_syntax) in (parent_path flat_names thy', defs', eqns @ [eqns'], rep_congs @ [cong'], dist_lemmas @ [dist]) end; - val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = foldl dt_constr_defs + val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs ((thy3 |> Theory.add_consts_i iso_decls |> parent_path flat_names, [], [], [], []), hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax); @@ -300,7 +300,7 @@ val newT_iso_inj_thms = map prove_newT_iso_inj_thm (new_type_names ~~ newT_iso_axms ~~ newTs ~~ - take (length newTs, rep_set_names)); + Library.take (length newTs, rep_set_names)); (********* isomorphisms between existing types and "unfolded" types *******) @@ -318,8 +318,8 @@ fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) = let val argTs = map (typ_of_dtyp descr' sorts) cargs; - val T = nth_elem (k, recTs); - val rep_name = nth_elem (k, all_rep_names); + val T = List.nth (recTs, k); + val rep_name = List.nth (all_rep_names, k); val rep_const = Const (rep_name, T --> Univ_elT); val constr = Const (cname, argTs ---> T); @@ -334,17 +334,17 @@ Ts @ [Us ---> Univ_elT]) else (i2 + 1, i2', ts @ [mk_lim (Us, - Const (nth_elem (j, all_rep_names), U --> Univ_elT) $ + Const (List.nth (all_rep_names, j), U --> Univ_elT) $ app_bnds (mk_Free "x" T' i2) (length Us))], Ts) | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts)) end; - val (i2, i2', ts, Ts) = foldl (process_arg ks) ((1, 1, [], []), cargs); + val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs); val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1))); val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1))); val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i); - val (_, _, ts', _) = foldl (process_arg []) ((1, 1, [], []), cargs); + val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs); val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i)) @@ -356,16 +356,16 @@ let val ks = map fst ds; val (_, (tname, _, _)) = hd ds; - val {rec_rewrites, rec_names, ...} = the (Symtab.lookup (dt_info, tname)); + val {rec_rewrites, rec_names, ...} = valOf (Symtab.lookup (dt_info, tname)); fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) = let - val (fs', eqns', _) = foldl (make_iso_def k ks (length constrs)) + val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs)) ((fs, eqns, 1), constrs); - val iso = (nth_elem (k, recTs), nth_elem (k, all_rep_names)) + val iso = (List.nth (recTs, k), List.nth (all_rep_names, k)) in (fs', eqns', isos @ [iso]) end; - val (fs, eqns, isos) = foldl process_dt (([], [], []), ds); + val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds); val fTs = map fastype_of fs; val defs = map (fn (rec_name, (T, iso_name)) => ((Sign.base_name iso_name) ^ "_def", equals (T --> Univ_elT) $ Const (iso_name, T --> Univ_elT) $ @@ -380,7 +380,7 @@ in (thy', char_thms' @ char_thms) end; - val (thy5, iso_char_thms) = foldr make_iso_defs + val (thy5, iso_char_thms) = Library.foldr make_iso_defs (tl descr, (add_path flat_names big_name thy4, [])); (* prove isomorphism properties *) @@ -412,13 +412,13 @@ fun prove_iso_thms (ds, (inj_thms, elem_thms)) = let val (_, (tname, _, _)) = hd ds; - val {induction, ...} = the (Symtab.lookup (dt_info, tname)); + val {induction, ...} = valOf (Symtab.lookup (dt_info, tname)); fun mk_ind_concl (i, _) = let - val T = nth_elem (i, recTs); - val Rep_t = Const (nth_elem (i, all_rep_names), T --> Univ_elT); - val rep_set_name = nth_elem (i, rep_set_names) + val T = List.nth (recTs, i); + val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT); + val rep_set_name = List.nth (rep_set_names, i) in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $ HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $ HOLogic.mk_eq (mk_Free "x" T i, Bound 0)), @@ -469,7 +469,7 @@ in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm)) end; - val (iso_inj_thms_unfolded, iso_elem_thms) = foldr prove_iso_thms + val (iso_inj_thms_unfolded, iso_elem_thms) = Library.foldr prove_iso_thms (tl descr, ([], map #3 newT_iso_axms)); val iso_inj_thms = map snd newT_iso_inj_thms @ iso_inj_thms_unfolded; @@ -494,7 +494,7 @@ iso_inj_thms_unfolded; val iso_thms = if length descr = 1 then [] else - drop (length newTs, split_conj_thm + Library.drop (length newTs, split_conj_thm (prove_goalw_cterm [] (cterm_of (Theory.sign_of thy5) iso_t) (fn _ => [(indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1, REPEAT (rtac TrueI 1), @@ -503,7 +503,7 @@ rewrite_goals_tac (map symmetric range_eqs), REPEAT (EVERY [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @ - flat (map (mk_funs_inv o #1) newT_iso_axms)) 1), + List.concat (map (mk_funs_inv o #1) newT_iso_axms)) 1), TRY (hyp_subst_tac 1), rtac (sym RS range_eqI) 1, resolve_tac iso_char_thms 1])]))); @@ -513,7 +513,7 @@ map2 (fn (r_inj, r) => f_myinv_f OF [r_inj, r RS mp]) (iso_inj_thms_unfolded, iso_thms); - val Abs_inverse_thms = flat (map mk_funs_inv Abs_inverse_thms'); + val Abs_inverse_thms = List.concat (map mk_funs_inv Abs_inverse_thms'); (******************* freeness theorems for constructors *******************) @@ -596,23 +596,23 @@ fun mk_indrule_lemma ((prems, concls), ((i, _), T)) = let - val Rep_t = Const (nth_elem (i, all_rep_names), T --> Univ_elT) $ + val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $ mk_Free "x" T i; val Abs_t = if i < length newTs then Const (Sign.intern_const (Theory.sign_of thy6) - ("Abs_" ^ (nth_elem (i, new_type_names))), Univ_elT --> T) + ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T) else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $ - Const (nth_elem (i, all_rep_names), T --> Univ_elT) + Const (List.nth (all_rep_names, i), T --> Univ_elT) in (prems @ [HOLogic.imp $ HOLogic.mk_mem (Rep_t, - Const (nth_elem (i, rep_set_names), UnivT)) $ + Const (List.nth (rep_set_names, i), UnivT)) $ (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) end; val (indrule_lemma_prems, indrule_lemma_concls) = - foldl mk_indrule_lemma (([], []), (descr' ~~ recTs)); + Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs)); val cert = cterm_of (Theory.sign_of thy6); diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/inductive_codegen.ML --- a/src/HOL/Tools/inductive_codegen.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/inductive_codegen.ML Thu Mar 03 12:43:01 2005 +0100 @@ -49,12 +49,12 @@ in (case concl_of thm of _ $ (Const ("op :", _) $ _ $ t) => (case head_of t of Const (s, _) => - let val cs = foldr add_term_consts (prems_of thm, []) + let val cs = Library.foldr add_term_consts (prems_of thm, []) in (CodegenData.put {intros = Symtab.update ((s, - if_none (Symtab.lookup (intros, s)) [] @ [thm]), intros), - graph = foldr (uncurry (Graph.add_edge o pair s)) - (cs, foldl add_node (graph, s :: cs)), + getOpt (Symtab.lookup (intros, s), []) @ [thm]), intros), + graph = Library.foldr (uncurry (Graph.add_edge o pair s)) + (cs, Library.foldl add_node (graph, s :: cs)), eqns = eqns} thy, thm) end | _ => (warn thm; p)) @@ -62,7 +62,7 @@ Const (s, _) => (CodegenData.put {intros = intros, graph = graph, eqns = Symtab.update ((s, - if_none (Symtab.lookup (eqns, s)) [] @ [thm]), eqns)} thy, thm) + getOpt (Symtab.lookup (eqns, s), []) @ [thm]), eqns)} thy, thm) | _ => (warn thm; p)) | _ => (warn thm; p)) end; @@ -77,7 +77,7 @@ let val SOME names = find_first (fn xs => s mem xs) (Graph.strong_conn graph) in SOME (names, preprocess thy - (flat (map (fn s => the (Symtab.lookup (intros, s))) names))) + (List.concat (map (fn s => valOf (Symtab.lookup (intros, s))) names))) end end; @@ -117,14 +117,14 @@ in (case (optf, strip_comb t) of (SOME f, (Const (name, _), args)) => (case assoc (extra_fs, name) of - NONE => overwrite (fs, (name, if_none - (apsome (mg_factor f) (assoc (fs, name))) f)) + NONE => overwrite (fs, (name, getOpt + (Option.map (mg_factor f) (assoc (fs, name)), f))) | SOME (fs', f') => (mg_factor f (FFix f'); - foldl (infer_factors sg extra_fs) - (fs, map (apsome FFix) fs' ~~ args))) + Library.foldl (infer_factors sg extra_fs) + (fs, map (Option.map FFix) fs' ~~ args))) | (SOME f, (Var ((name, _), _), [])) => - overwrite (fs, (name, if_none - (apsome (mg_factor f) (assoc (fs, name))) f)) + overwrite (fs, (name, getOpt + (Option.map (mg_factor f) (assoc (fs, name)), f))) | (NONE, _) => fs | _ => err "Illegal term") handle Factors => err "Product factor mismatch in" @@ -139,7 +139,7 @@ fun is_constrt thy = let - val cnstrs = flat (flat (map + val cnstrs = List.concat (List.concat (map (map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd) (Symtab.dest (DatatypePackage.get_datatypes thy)))); fun check t = (case strip_comb t of @@ -167,11 +167,11 @@ string_of_mode ms)) modes)); val term_vs = map (fst o fst o dest_Var) o term_vars; -val terms_vs = distinct o flat o (map term_vs); +val terms_vs = distinct o List.concat o (map term_vs); (** collect all Vars in a term (with duplicates!) **) fun term_vTs t = map (apfst fst o dest_Var) - (filter is_Var (foldl_aterms (op :: o Library.swap) ([], t))); + (List.filter is_Var (foldl_aterms (op :: o Library.swap) ([], t))); fun get_args _ _ [] = ([], []) | get_args is i (x::xs) = (if i mem is then apfst else apsnd) (cons x) @@ -190,18 +190,18 @@ fun cprod ([], ys) = [] | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys); -fun cprods xss = foldr (map op :: o cprod) (xss, [[]]); +fun cprods xss = Library.foldr (map op :: o cprod) (xss, [[]]); datatype mode = Mode of (int list option list * int list) * mode option list; fun modes_of modes t = let - fun mk_modes name args = flat + fun mk_modes name args = List.concat (map (fn (m as (iss, is)) => map (Mode o pair m) (cprods (map (fn (NONE, _) => [NONE] | (SOME js, arg) => map SOME - (filter (fn Mode ((_, js'), _) => js=js') (modes_of modes arg))) - (iss ~~ args)))) (the (assoc (modes, name)))) + (List.filter (fn Mode ((_, js'), _) => js=js') (modes_of modes arg))) + (iss ~~ args)))) (valOf (assoc (modes, name)))) in (case strip_comb t of (Const ("op =", Type (_, [T, _])), _) => @@ -215,14 +215,14 @@ datatype indprem = Prem of term list * term | Sidecond of term; fun select_mode_prem thy modes vs ps = - find_first (is_some o snd) (ps ~~ map + find_first (isSome o snd) (ps ~~ map (fn Prem (us, t) => find_first (fn Mode ((_, is), _) => let val (in_ts, out_ts) = get_args is 1 us; - val (out_ts', in_ts') = partition (is_constrt thy) out_ts; - val vTs = flat (map term_vTs out_ts'); + val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts; + val vTs = List.concat (map term_vTs out_ts'); val dupTs = map snd (duplicates vTs) @ - mapfilter (curry assoc vTs) vs; + List.mapPartial (curry assoc vTs) vs; in terms_vs (in_ts @ in_ts') subset vs andalso forall (is_eqT o fastype_of) in_ts' andalso @@ -235,7 +235,7 @@ fun check_mode_clause thy arg_vs modes (iss, is) (ts, ps) = let - val modes' = modes @ mapfilter + val modes' = modes @ List.mapPartial (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)])) (arg_vs ~~ iss); fun check_mode_prems vs [] = SOME vs @@ -244,11 +244,11 @@ | SOME (x, _) => check_mode_prems (case x of Prem (us, _) => vs union terms_vs us | _ => vs) (filter_out (equal x) ps)); - val (in_ts, in_ts') = partition (is_constrt thy) (fst (get_args is 1 ts)); + val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (get_args is 1 ts)); val in_vs = terms_vs in_ts; val concl_vs = terms_vs ts in - forall is_eqT (map snd (duplicates (flat (map term_vTs in_ts)))) andalso + forall is_eqT (map snd (duplicates (List.concat (map term_vTs in_ts)))) andalso forall (is_eqT o fastype_of) in_ts' andalso (case check_mode_prems (arg_vs union in_vs) ps of NONE => false @@ -257,7 +257,7 @@ fun check_modes_pred thy arg_vs preds modes (p, ms) = let val SOME rs = assoc (preds, p) - in (p, filter (fn m => case find_index + in (p, List.filter (fn m => case find_index (not o check_mode_clause thy arg_vs modes m) rs of ~1 => true | i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^ @@ -283,7 +283,7 @@ in mk_eqs x xs end; fun mk_tuple xs = Pretty.block (Pretty.str "(" :: - flat (separate [Pretty.str ",", Pretty.brk 1] (map single xs)) @ + List.concat (separate [Pretty.str ",", Pretty.brk 1] (map single xs)) @ [Pretty.str ")"]); (* convert nested pairs to n-tuple *) @@ -351,8 +351,8 @@ | is_exhaustive _ = false; fun compile_match nvs eq_ps out_ps success_p can_fail = - let val eqs = flat (separate [Pretty.str " andalso", Pretty.brk 1] - (map single (flat (map (mk_eq o snd) nvs) @ eq_ps))); + let val eqs = List.concat (separate [Pretty.str " andalso", Pretty.brk 1] + (map single (List.concat (map (mk_eq o snd) nvs) @ eq_ps))); in Pretty.block ([Pretty.str "(fn ", mk_tuple out_ps, Pretty.str " =>", Pretty.brk 1] @ @@ -367,7 +367,7 @@ fun modename thy s (iss, is) = space_implode "__" (mk_const_id (sign_of thy) s :: - map (space_implode "_" o map string_of_int) (mapfilter I iss @ [is])); + map (space_implode "_" o map string_of_int) (List.mapPartial I iss @ [is])); fun compile_expr thy dep brack (gr, (NONE, t)) = apsnd single (invoke_codegen thy dep brack (gr, t)) @@ -380,13 +380,13 @@ (compile_expr thy dep true) (gr, ms ~~ args); in (gr', (if brack andalso not (null ps) then single o parens o Pretty.block else I) - (flat (separate [Pretty.brk 1] + (List.concat (separate [Pretty.brk 1] ([Pretty.str (modename thy name mode)] :: ps)))) end; fun compile_clause thy gr dep all_vs arg_vs modes (iss, is) (ts, ps) = let - val modes' = modes @ mapfilter + val modes' = modes @ List.mapPartial (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)])) (arg_vs ~~ iss); @@ -404,7 +404,7 @@ foldl_map check_constrt ((all_vs, []), in_ts); fun is_ind t = (case head_of t of - Const (s, _) => s = "op =" orelse is_some (assoc (modes, s)) + Const (s, _) => s = "op =" orelse isSome (assoc (modes, s)) | Var ((s, _), _) => s mem arg_vs); fun compile_prems out_ts' vs names gr [] = @@ -426,7 +426,7 @@ end | compile_prems out_ts vs names gr ps = let - val vs' = distinct (flat (vs :: map term_vs out_ts)); + val vs' = distinct (List.concat (vs :: map term_vs out_ts)); val SOME (p, mode as SOME (Mode ((_, js), _))) = select_mode_prem thy modes' vs' ps; val ps' = filter_out (equal p) ps; @@ -484,16 +484,16 @@ (Pretty.str (prfx ^ modename thy s mode) :: map Pretty.str arg_vs) @ [Pretty.str " inp ="]), Pretty.brk 1] @ - flat (separate [Pretty.str " ++", Pretty.brk 1] (map single cl_ps)))) + List.concat (separate [Pretty.str " ++", Pretty.brk 1] (map single cl_ps)))) end; fun compile_preds thy gr dep all_vs arg_vs modes preds = let val ((gr', _), prs) = foldl_map (fn ((gr, prfx), (s, cls)) => foldl_map (fn ((gr', prfx'), mode) => compile_pred thy gr' dep prfx' all_vs arg_vs modes s cls mode) - ((gr, prfx), the (assoc (modes, s)))) ((gr, "fun "), preds) + ((gr, prfx), valOf (assoc (modes, s)))) ((gr, "fun "), preds) in - (gr', space_implode "\n\n" (map Pretty.string_of (flat prs)) ^ ";\n\n") + (gr', space_implode "\n\n" (map Pretty.string_of (List.concat prs)) ^ ";\n\n") end; (**** processing of introduction rules ****) @@ -502,7 +502,7 @@ (string * (int list option list * int list) list) list * (string * (int list list option list * int list list)) list; -fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip +fun lookup_modes gr dep = apfst List.concat (apsnd List.concat (ListPair.unzip (map ((fn (SOME (Modes x), _) => x | _ => ([], [])) o Graph.get_node gr) (Graph.all_preds gr [dep])))); @@ -520,13 +520,13 @@ | SOME xs' => xs inter xs') :: constrain cs ys; fun mk_extra_defs thy gr dep names ts = - foldl (fn (gr, name) => + Library.foldl (fn (gr, name) => if name mem names then gr else (case get_clauses thy name of NONE => gr | SOME (names, intrs) => mk_ind_def thy gr dep names [] [] (prep_intrs intrs))) - (gr, foldr add_term_consts (ts, [])) + (gr, Library.foldr add_term_consts (ts, [])) and mk_ind_def thy gr dep names modecs factorcs intrs = let val ids = map (mk_const_id (sign_of thy)) names @@ -534,7 +534,7 @@ let val _ $ (_ $ _ $ u) = Logic.strip_imp_concl (hd intrs); val (_, args) = strip_comb u; - val arg_vs = flat (map term_vs args); + val arg_vs = List.concat (map term_vs args); fun dest_prem factors (_ $ (p as (Const ("op :", _) $ t $ u))) = (case assoc (factors, case head_of u of @@ -551,11 +551,11 @@ val Const (name, _) = head_of u; val prems = map (dest_prem factors) (Logic.strip_imp_prems intr); in - (overwrite (clauses, (name, if_none (assoc (clauses, name)) [] @ - [(split_prod [] (the (assoc (factors, name))) t, prems)]))) + (overwrite (clauses, (name, getOpt (assoc (clauses, name), []) @ + [(split_prod [] (valOf (assoc (factors, name))) t, prems)]))) end; - fun check_set (Const (s, _)) = s mem names orelse is_some (get_clauses thy s) + fun check_set (Const (s, _)) = s mem names orelse isSome (get_clauses thy s) | check_set (Var ((s, _), _)) = s mem arg_vs | check_set _ = false; @@ -571,13 +571,13 @@ (Graph.new_node (hd ids, (NONE, "")) gr)) (hd ids) names intrs; val (extra_modes, extra_factors) = lookup_modes gr' (hd ids); val fs = constrain factorcs (map (apsnd dest_factors) - (foldl (add_prod_factors extra_factors) ([], flat (map (fn t => + (Library.foldl (add_prod_factors extra_factors) ([], List.concat (map (fn t => Logic.strip_imp_concl t :: Logic.strip_imp_prems t) intrs)))); - val factors = mapfilter (fn (name, f) => + val factors = List.mapPartial (fn (name, f) => if name mem arg_vs then NONE else SOME (name, (map (curry assoc fs) arg_vs, f))) fs; val clauses = - foldl (add_clause (fs @ map (apsnd snd) extra_factors)) ([], intrs); + Library.foldl (add_clause (fs @ map (apsnd snd) extra_factors)) ([], intrs); val modes = constrain modecs (infer_modes thy extra_modes factors arg_vs clauses); val _ = print_factors factors; @@ -606,9 +606,9 @@ val gr1 = mk_extra_defs thy (mk_ind_def thy gr dep names [] [] (prep_intrs intrs)) dep names [u]; val (modes, factors) = lookup_modes gr1 dep; - val ts = split_prod [] (snd (the (assoc (factors, s)))) t; + val ts = split_prod [] (snd (valOf (assoc (factors, s)))) t; val (ts', is) = if is_query then - fst (foldl mk_mode ((([], []), 1), ts)) + fst (Library.foldl mk_mode ((([], []), 1), ts)) else (ts, 1 upto length ts); val mode = find_mode s u modes is; val (gr2, in_ps) = foldl_map @@ -635,7 +635,7 @@ SOME (gr2, (if brack then parens else I) (Pretty.block ([Pretty.str "Seq.list_of", Pretty.brk 1, Pretty.str "("] @ - conv_ntuple' (snd (the (assoc (factors, s)))) + conv_ntuple' (snd (valOf (assoc (factors, s)))) (HOLogic.dest_setT (fastype_of u)) (ps @ [Pretty.brk 1, Pretty.str "()"]) @ [Pretty.str ")"]))) @@ -718,7 +718,7 @@ fun ?? b = if b then Seq.single () else Seq.empty; -fun ?! s = is_some (Seq.pull s); +fun ?! s = isSome (Seq.pull s); fun op_61__1 x = Seq.single x; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/inductive_package.ML --- a/src/HOL/Tools/inductive_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/inductive_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -124,7 +124,7 @@ fun put_inductives names info thy = let fun upd ((tab, monos), name) = (Symtab.update_new ((name, info), tab), monos); - val tab_monos = foldl upd (InductiveData.get thy, names) + val tab_monos = Library.foldl upd (InductiveData.get thy, names) handle Symtab.DUP name => error ("Duplicate definition of (co)inductive set " ^ quote name); in InductiveData.put tab_monos thy end; @@ -183,16 +183,16 @@ fun varify (t, (i, ts)) = let val t' = map_term_types (incr_tvar (i + 1)) (#1 (Type.varify (t, []))) in (maxidx_of_term t', t'::ts) end; - val (i, cs') = foldr varify (cs, (~1, [])); - val (i', intr_ts') = foldr varify (intr_ts, (i, [])); - val rec_consts = foldl add_term_consts_2 ([], cs'); - val intr_consts = foldl add_term_consts_2 ([], intr_ts'); + val (i, cs') = Library.foldr varify (cs, (~1, [])); + val (i', intr_ts') = Library.foldr varify (intr_ts, (i, [])); + val rec_consts = Library.foldl add_term_consts_2 ([], cs'); + val intr_consts = Library.foldl add_term_consts_2 ([], intr_ts'); fun unify (env, (cname, cT)) = - let val consts = map snd (filter (fn c => fst c = cname) intr_consts) - in foldl (fn ((env', j'), Tp) => (Type.unify tsig (env', j') Tp)) + let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts) + in Library.foldl (fn ((env', j'), Tp) => (Type.unify tsig (env', j') Tp)) (env, (replicate (length consts) cT) ~~ consts) end; - val (env, _) = foldl unify ((Vartab.empty, i'), rec_consts); + val (env, _) = Library.foldl unify ((Vartab.empty, i'), rec_consts); fun typ_subst_TVars_2 env T = let val T' = typ_subst_TVars_Vartab env T in if T = T' then T else typ_subst_TVars_2 env T' end; val subst = fst o Type.freeze_thaw o @@ -237,7 +237,7 @@ fun mg_prod_factors ts (fs, t $ u) = if t mem ts then let val f = prod_factors [] u - in overwrite (fs, (t, f inter if_none (assoc (fs, t)) f)) end + in overwrite (fs, (t, f inter (curry getOpt) (assoc (fs, t)) f)) end else mg_prod_factors ts (mg_prod_factors ts (fs, t), u) | mg_prod_factors ts (fs, Abs (_, _, t)) = mg_prod_factors ts (fs, t) | mg_prod_factors ts (fs, _) = fs; @@ -256,7 +256,7 @@ fun mk_tuple p ps (Type ("*", [T1, T2])) (tms as t::_) = if p mem ps then HOLogic.mk_prod (mk_tuple (1::p) ps T1 tms, - mk_tuple (2::p) ps T2 (drop (length (prodT_factors (1::p) ps T1), tms))) + mk_tuple (2::p) ps T2 (Library.drop (length (prodT_factors (1::p) ps T1), tms))) else t | mk_tuple _ _ _ (t::_) = t; @@ -271,12 +271,12 @@ val remove_split = rewrite_rule [split_conv RS eq_reflection]; -fun split_rule_vars vs rl = standard (remove_split (foldr split_rule_var' +fun split_rule_vars vs rl = standard (remove_split (Library.foldr split_rule_var' (mg_prod_factors vs ([], Thm.prop_of rl), rl))); -fun split_rule vs rl = standard (remove_split (foldr split_rule_var' - (mapfilter (fn (t as Var ((a, _), _)) => - apsome (pair t) (assoc (vs, a))) (term_vars (Thm.prop_of rl)), rl))); +fun split_rule vs rl = standard (remove_split (Library.foldr split_rule_var' + (List.mapPartial (fn (t as Var ((a, _), _)) => + Option.map (pair t) (assoc (vs, a))) (term_vars (Thm.prop_of rl)), rl))); (** process rules **) @@ -315,7 +315,7 @@ if u mem cs then if exists (Logic.occs o rpair t) cs then err_in_rule sg name rule "Recursion term on left of member symbol" - else seq check_prem (prems ~~ aprems) + else List.app check_prem (prems ~~ aprems) else err_in_rule sg name rule bad_concl | Const ("all", _) $ _ => err_in_rule sg name rule all_not_allowed | _ => err_in_rule sg name rule bad_concl); @@ -337,7 +337,7 @@ fun mk_elims cs cTs params intr_ts intr_names = let - val used = foldr add_term_names (intr_ts, []); + val used = Library.foldr add_term_names (intr_ts, []); val [aname, pname] = variantlist (["a", "P"], used); val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT)); @@ -353,9 +353,9 @@ val a = Free (aname, T); fun mk_elim_prem (_, t, ts) = - list_all_free (map dest_Free ((foldr add_term_frees (t::ts, [])) \\ params), + list_all_free (map dest_Free ((Library.foldr add_term_frees (t::ts, [])) \\ params), Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (a, t)) :: ts, P)); - val c_intrs = (filter (equal c o #1 o #1) intrs); + val c_intrs = (List.filter (equal c o #1 o #1) intrs); in (Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (a, c)) :: map mk_elim_prem (map #1 c_intrs), P), map #2 c_intrs) @@ -369,7 +369,7 @@ fun mk_indrule cs cTs params intr_ts = let - val used = foldr add_term_names (intr_ts, []); + val used = Library.foldr add_term_names (intr_ts, []); (* predicates for induction rule *) @@ -407,22 +407,22 @@ HOLogic.dest_Trueprop (Logic.strip_imp_concl r) in list_all_free (frees, - Logic.list_implies (map HOLogic.mk_Trueprop (foldr mk_prem + Logic.list_implies (map HOLogic.mk_Trueprop (Library.foldr mk_prem (map HOLogic.dest_Trueprop (Logic.strip_imp_prems r), [])), - HOLogic.mk_Trueprop (the (pred_of u) $ t))) + HOLogic.mk_Trueprop (valOf (pred_of u) $ t))) end; val ind_prems = map mk_ind_prem intr_ts; - val factors = foldl (mg_prod_factors preds) ([], ind_prems); + val factors = Library.foldl (mg_prod_factors preds) ([], ind_prems); (* make conclusions for induction rules *) fun mk_ind_concl ((c, P), (ts, x)) = let val T = HOLogic.dest_setT (fastype_of c); - val ps = if_none (assoc (factors, P)) []; + val ps = getOpt (assoc (factors, P), []); val Ts = prodT_factors [] ps T; - val (frees, x') = foldr (fn (T', (fs, s)) => + val (frees, x') = Library.foldr (fn (T', (fs, s)) => ((Free (s, T'))::fs, Symbol.bump_string s)) (Ts, ([], x)); val tuple = mk_tuple [] ps T frees; in ((HOLogic.mk_binop "op -->" @@ -430,7 +430,7 @@ end; val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj - (fst (foldr mk_ind_concl (cs ~~ preds, ([], "xa"))))) + (fst (Library.foldr mk_ind_concl (cs ~~ preds, ([], "xa"))))) in (preds, ind_prems, mutual_ind_concl, map (apfst (fst o dest_Free)) factors) @@ -481,7 +481,7 @@ Goals.prove_goalw_cterm [] (*NO quick_and_dirty_prove_goalw_cterm here!*) (Thm.cterm_of (Theory.sign_of thy) (HOLogic.mk_Trueprop (Const (mono_name, (setT --> setT) --> HOLogic.boolT) $ fp_fun))) - (fn _ => [rtac monoI 1, REPEAT (ares_tac (flat (map mk_mono monos) @ get_monos thy) 1)])); + (fn _ => [rtac monoI 1, REPEAT (ares_tac (List.concat (map mk_mono monos) @ get_monos thy) 1)])); (* prove introduction rules *) @@ -628,7 +628,7 @@ val dummy = if !trace then (writeln "ind_prems = "; - seq (writeln o Sign.string_of_term sign) ind_prems) + List.app (writeln o Sign.string_of_term sign) ind_prems) else (); (* make predicate for instantiation of abstract induction rule *) @@ -639,7 +639,7 @@ val Type (_, [T1, T2]) = T in Const ("Datatype.sum.sum_case", [T1 --> HOLogic.boolT, T2 --> HOLogic.boolT, T] ---> HOLogic.boolT) $ - mk_ind_pred T1 (take (n, Ps)) $ mk_ind_pred T2 (drop (n, Ps)) + mk_ind_pred T1 (Library.take (n, Ps)) $ mk_ind_pred T2 (Library.drop (n, Ps)) end; val ind_pred = mk_ind_pred sumT preds; @@ -655,7 +655,7 @@ (HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_vimage cs sumT (HOLogic.Collect_const sumT $ ind_pred) c, HOLogic.Collect_const (HOLogic.dest_setT (fastype_of c)) $ - nth_elem (find_index_eq c cs, preds))))) + List.nth (preds, find_index_eq c cs))))) (fn _ => [rtac vimage_Collect 1, rewrite_goals_tac sum_case_rewrites, rtac refl 1])) cs; val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct)); @@ -707,7 +707,7 @@ val fp_name = if coind then gfp_name else lfp_name; - val used = foldr add_term_names (intr_ts, []); + val used = Library.foldr add_term_names (intr_ts, []); val [sname, xname] = variantlist (["S", "x"], used); (* transform an introduction rule into a conjunction *) @@ -723,7 +723,7 @@ val Const ("op :", _) $ t $ u = HOLogic.dest_Trueprop (Logic.strip_imp_concl r) - in foldr (fn ((x, T), P) => HOLogic.mk_exists (x, T, P)) + in Library.foldr (fn ((x, T), P) => HOLogic.mk_exists (x, T, P)) (frees, foldr1 HOLogic.mk_conj (((HOLogic.eq_const sumT) $ Free (xname, sumT) $ (mk_inj cs sumT u t)):: (map (subst o HOLogic.dest_Trueprop) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/inductive_realizer.ML Thu Mar 03 12:43:01 2005 +0100 @@ -41,7 +41,7 @@ | strip _ t = t; in strip (add_term_free_names (t, [])) t end; -fun relevant_vars prop = foldr (fn +fun relevant_vars prop = Library.foldr (fn (Var ((a, i), T), vs) => (case strip_type T of (_, Type (s, _)) => if s mem ["bool", "set"] then (a, T) :: vs else vs | _ => vs) @@ -198,13 +198,13 @@ fun indrule_realizer thy induct raw_induct rsets params vs rec_names rss intrs dummies = let val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)); - val premss = mapfilter (fn (s, rs) => if s mem rsets then - SOME (map (fn r => nth_elem (find_index_eq (prop_of r) (map prop_of intrs), - prems_of raw_induct)) rs) else NONE) rss; - val concls' = mapfilter (fn (s, _) => if s mem rsets then + val premss = List.mapPartial (fn (s, rs) => if s mem rsets then + SOME (map (fn r => List.nth (prems_of raw_induct, + find_index_eq (prop_of r) (map prop_of intrs))) rs) else NONE) rss; + val concls' = List.mapPartial (fn (s, _) => if s mem rsets then find_first (fn concl => s mem term_consts concl) concls else NONE) rss; - val fs = flat (snd (foldl_map (fn (intrs, (prems, dummy)) => + val fs = List.concat (snd (foldl_map (fn (intrs, (prems, dummy)) => let val (intrs1, intrs2) = splitAt (length prems, intrs); val fs = map (fn (rule, intr) => @@ -213,9 +213,9 @@ HOLogic.unitT --> body_type (fastype_of (hd fs))) :: fs else fs) end) (intrs, (premss ~~ dummies)))); - val frees = foldl Term.add_frees ([], fs); + val frees = Library.foldl Term.add_frees ([], fs); val Ts = map fastype_of fs; - val rlzs = mapfilter (fn (a, concl) => + val rlzs = List.mapPartial (fn (a, concl) => let val T = Extraction.etype_of thy vs [] concl in if T = Extraction.nullT then NONE else SOME (list_comb (Const (a, Ts ---> T), fs)) @@ -225,8 +225,8 @@ val r = foldr1 HOLogic.mk_prod rlzs; val x = Free ("x", Extraction.etype_of thy vs [] (hd (prems_of induct))); fun name_of_fn intr = "r" ^ Sign.base_name (Thm.name_of_thm intr); - val r' = list_abs_free (mapfilter (fn intr => - apsome (pair (name_of_fn intr)) (assoc (frees, name_of_fn intr))) intrs, + val r' = list_abs_free (List.mapPartial (fn intr => + Option.map (pair (name_of_fn intr)) (assoc (frees, name_of_fn intr))) intrs, if length concls = 1 then r $ x else r) in if length concls = 1 then lambda x r' else r' @@ -253,7 +253,7 @@ val xs = rev (Term.add_vars ([], prop_of rule)); val vs1 = map Var (filter_out (fn ((a, _), _) => a mem rvs) xs); val rlzvs = rev (Term.add_vars ([], prop_of rrule)); - val vs2 = map (fn (ixn, _) => Var (ixn, the (assoc (rlzvs, ixn)))) xs; + val vs2 = map (fn (ixn, _) => Var (ixn, valOf (assoc (rlzvs, ixn)))) xs; val rs = gen_rems (op = o pairself fst) (rlzvs, xs); fun mk_prf _ [] prf = prf @@ -265,8 +265,8 @@ in (Thm.name_of_thm rule, (vs, if rt = Extraction.nullt then rt else - foldr (uncurry lambda) (vs1, rt), - foldr forall_intr_prf (vs2, mk_prf rs prems (Proofterm.proof_combP + Library.foldr (uncurry lambda) (vs1, rt), + Library.foldr forall_intr_prf (vs2, mk_prf rs prems (Proofterm.proof_combP (prf_of rrule, map PBound (length prems - 1 downto 0)))))) end; @@ -274,7 +274,7 @@ let val _ $ (_ $ _ $ S) = concl_of r; val (Const (s, _), _) = strip_comb S; - val rs = if_none (assoc (rss, s)) []; + val rs = getOpt (assoc (rss, s), []); in overwrite (rss, (s, rs @ [r])) end; fun add_ind_realizer rsets intrs induct raw_induct elims (thy, vs) = @@ -284,7 +284,7 @@ val (_ $ (_ $ _ $ S)) = Logic.strip_imp_concl (prop_of (hd intrs)); val (_, params) = strip_comb S; val params' = map dest_Var params; - val rss = foldl add_rule ([], intrs); + val rss = Library.foldl add_rule ([], intrs); val (prfx, _) = split_last (NameSpace.unpack (fst (hd rss))); val tnames = map (fn s => space_implode "_" (s ^ "T" :: vs)) rsets; val {path, ...} = Sign.rep_sg (sign_of thy); @@ -300,7 +300,7 @@ Theory.add_arities_i (map (fn s => (s, replicate ar HOLogic.typeS, HOLogic.typeS)) tnames) |> Extraction.add_typeof_eqns_i ty_eqs; - val dts = mapfilter (fn (s, rs) => if s mem rsets then + val dts = List.mapPartial (fn (s, rs) => if s mem rsets then SOME (dt_of_intrs thy1' vs rs) else NONE) rss; (** datatype representing computational content of inductive set **) @@ -311,7 +311,7 @@ (map #2 dts)) (map (pair false) dts) []) |>> Extraction.add_typeof_eqns_i ty_eqs |>> Extraction.add_realizes_eqns_i rlz_eqs; - fun get f x = if_none (apsome f x) []; + fun get f x = getOpt (Option.map f x, []); val rec_names = distinct (map (fst o dest_Const o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) (get #rec_thms dt_info)); val (_, constrss) = foldl_map (fn ((recs, dummies), (s, rs)) => @@ -328,7 +328,7 @@ (Extraction.realizes_of thy2 vs c (prop_of (forall_intr_list (map (cterm_of (sign_of thy2) o Var) (rev (Term.add_vars ([], prop_of intr)) \\ params')) intr)))) - (intrs ~~ flat constrss); + (intrs ~~ List.concat constrss); val rlzsets = distinct (map (fn rintr => snd (HOLogic.dest_mem (HOLogic.dest_Trueprop (Logic.strip_assums_concl rintr)))) rintrs); @@ -345,7 +345,7 @@ (** realizer for induction rule **) - val Ps = mapfilter (fn _ $ M $ P => if set_of M mem rsets then + val Ps = List.mapPartial (fn _ $ M $ P => if set_of M mem rsets then SOME (fst (fst (dest_Var (head_of P)))) else NONE) (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct))); @@ -383,13 +383,13 @@ let val (prem :: prems) = prems_of elim; fun reorder1 (p, intr) = - foldl (fn (t, ((s, _), T)) => all T $ lambda (Free (s, T)) t) + Library.foldl (fn (t, ((s, _), T)) => all T $ lambda (Free (s, T)) t) (strip_all p, Term.add_vars ([], prop_of intr) \\ params'); fun reorder2 (intr, i) = let val fs1 = term_vars (prop_of intr) \\ params; val fs2 = Term.add_vars ([], prop_of intr) \\ params' - in foldl (fn (t, x) => lambda (Var x) t) + in Library.foldl (fn (t, x) => lambda (Var x) t) (list_comb (Bound (i + length fs1), fs1), fs2) end; val p = Logic.list_implies @@ -423,24 +423,24 @@ (** add realizers to theory **) - val rintr_thms = flat (map (fn (_, rs) => map (fn r => nth_elem - (find_index_eq r intrs, #intrs ind_info)) rs) rss); - val thy4 = foldl add_ind_realizer (thy3, subsets Ps); + val rintr_thms = List.concat (map (fn (_, rs) => map (fn r => List.nth + (#intrs ind_info, find_index_eq r intrs)) rs) rss); + val thy4 = Library.foldl add_ind_realizer (thy3, subsets Ps); val thy5 = Extraction.add_realizers_i (map (mk_realizer thy4 vs params') (map (fn ((rule, rrule), c) => ((rule, rrule), list_comb (c, map Var (rev (Term.add_vars ([], prop_of rule)) \\ params')))) - (flat (map snd rss) ~~ rintr_thms ~~ flat constrss))) thy4; - val elimps = mapfilter (fn (s, intrs) => if s mem rsets then - apsome (rpair intrs) (find_first (fn (thm, _) => + (List.concat (map snd rss) ~~ rintr_thms ~~ List.concat constrss))) thy4; + val elimps = List.mapPartial (fn (s, intrs) => if s mem rsets then + Option.map (rpair intrs) (find_first (fn (thm, _) => s mem term_consts (hd (prems_of thm))) (elims ~~ #elims ind_info)) else NONE) rss; - val thy6 = foldl (fn (thy, p as (((((elim, _), _), _), _), _)) => thy |> + val thy6 = Library.foldl (fn (thy, p as (((((elim, _), _), _), _), _)) => thy |> add_elim_realizer [] p |> add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (concl_of elim))))] p) (thy5, elimps ~~ get #case_thms dt_info ~~ case_names ~~ dummies) - in Theory.add_path (NameSpace.pack (if_none path [])) thy6 end; + in Theory.add_path (NameSpace.pack (getOpt (path, []))) thy6 end; fun add_ind_realizers name rsets thy = let @@ -452,7 +452,7 @@ val vss = sort (int_ord o pairself length) (subsets (map fst (relevant_vars S))) in - foldl (add_ind_realizer rsets intrs induct raw_induct elims) (thy, vss) + Library.foldl (add_ind_realizer rsets intrs induct raw_induct elims) (thy, vss) end fun rlz_attrib arg (thy, thm) = @@ -462,7 +462,7 @@ (case HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of thm)) of [_] => [set_of (HOLogic.dest_Trueprop (hd (prems_of thm)))] | xs => map (set_of o fst o HOLogic.dest_imp) xs) - handle TERM _ => err () | LIST _ => err (); + handle TERM _ => err () | Empty => err (); in (add_ind_realizers (hd sets) (case arg of NONE => sets | SOME NONE => [] diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/meson.ML --- a/src/HOL/Tools/meson.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/meson.ML Thu Mar 03 12:43:01 2005 +0100 @@ -164,7 +164,7 @@ fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); (*TAUTOLOGY CHECK SHOULD NOT BE NECESSARY!*) - fun sort_clauses ths = sort (make_ord fewerlits) (filter (not o is_taut) ths); + fun sort_clauses ths = sort (make_ord fewerlits) (List.filter (not o is_taut) ths); (*Convert all suitable free variables to schematic variables*) fun generalize th = forall_elim_vars 0 (forall_intr_frees th); @@ -188,12 +188,12 @@ let fun name1 (th, (k,ths)) = (k-1, Thm.name_thm (label ^ string_of_int k, th) :: ths) - in fn ths => #2 (foldr name1 (ths, (length ths, []))) end; + in fn ths => #2 (Library.foldr name1 (ths, (length ths, []))) end; (*Find an all-negative support clause*) fun is_negative th = forall (not o #1) (literals (prop_of th)); - val neg_clauses = filter is_negative; + val neg_clauses = List.filter is_negative; (***** MESON PROOF PROCEDURE *****) @@ -211,7 +211,7 @@ fun has_reps [] = false | has_reps [_] = false | has_reps [t,u] = (t aconv u) - | has_reps ts = (foldl ins_term (Net.empty, ts); false) + | has_reps ts = (Library.foldl ins_term (Net.empty, ts); false) handle INSERT => true; (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) @@ -239,7 +239,7 @@ (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) local fun addconcl(prem,sz) = size_of_term(Logic.strip_assums_concl prem) + sz in -fun size_of_subgoals st = foldr addconcl (prems_of st, 0) +fun size_of_subgoals st = Library.foldr addconcl (prems_of st, 0) end; (*Negation Normal Form*) @@ -265,12 +265,12 @@ (*Make clauses from a list of theorems, previously Skolemized and put into nnf. The resulting clauses are HOL disjunctions.*) fun make_clauses ths = - sort_clauses (map (generalize o nodups) (foldr cnf (ths,[]))); + sort_clauses (map (generalize o nodups) (Library.foldr cnf (ths,[]))); (*Convert a list of clauses to (contrapositive) Horn clauses*) fun make_horns ths = name_thms "Horn#" - (gen_distinct Drule.eq_thm_prop (foldr (add_contras clause_rules) (ths,[]))); + (gen_distinct Drule.eq_thm_prop (Library.foldr (add_contras clause_rules) (ths,[]))); (*Could simply use nprems_of, which would count remaining subgoals -- no discrimination as to their size! With BEST_FIRST, fails for problem 41.*) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/primrec_package.ML --- a/src/HOL/Tools/primrec_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/primrec_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -76,7 +76,7 @@ NONE => (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns | SOME (_, rpos', eqns) => - if is_some (assoc (eqns, cname)) then + if isSome (assoc (eqns, cname)) then raise RecError "constructor already occurred as pattern" else if rpos <> rpos' then raise RecError "position of recursive argument inconsistent" @@ -90,7 +90,7 @@ fun process_fun sign descr rec_eqns ((i, fname), (fnames, fnss)) = let - val (_, (tname, _, constrs)) = nth_elem (i, descr); + val (_, (tname, _, constrs)) = List.nth (descr, i); (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) @@ -104,11 +104,11 @@ if is_Const f andalso (fst (dest_Const f)) mem (map fst rec_eqns) then let val (fname', _) = dest_Const f; - val (_, rpos, _) = the (assoc (rec_eqns, fname')); - val ls = take (rpos, ts); - val rest = drop (rpos, ts); + val (_, rpos, _) = valOf (assoc (rec_eqns, fname')); + val ls = Library.take (rpos, ts); + val rest = Library.drop (rpos, ts); val (x', rs) = (hd rest, tl rest) - handle LIST _ => raise RecError ("not enough arguments\ + handle Empty => raise RecError ("not enough arguments\ \ in recursive application\nof function " ^ quote fname' ^ " on rhs"); val (x, xs) = strip_comb x' in @@ -140,7 +140,7 @@ (fnames', fnss', (Const ("arbitrary", dummyT))::fns)) | SOME (ls, cargs', rs, rhs, eq) => let - val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); + val recs = List.filter (is_rec_type o snd) (cargs' ~~ cargs); val rargs = map fst recs; val subs = map (rpair dummyT o fst) (rev (rename_wrt_term rhs rargs)); @@ -160,8 +160,8 @@ raise RecError ("inconsistent functions for datatype " ^ quote tname) else let - val (_, _, eqns) = the (assoc (rec_eqns, fname)); - val (fnames', fnss', fns) = foldr (trans eqns) + val (_, _, eqns) = valOf (assoc (rec_eqns, fname)); + val (fnames', fnss', fns) = Library.foldr (trans eqns) (constrs, ((i, fname)::fnames, fnss, [])) in (fnames', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') @@ -179,7 +179,7 @@ NONE => let val dummy_fns = map (fn (_, cargs) => Const ("arbitrary", - replicate ((length cargs) + (length (filter is_rec_type cargs))) + replicate ((length cargs) + (length (List.filter is_rec_type cargs))) dummyT ---> HOLogic.unitT)) constrs; val _ = warning ("No function definition for datatype " ^ quote tname) in @@ -192,7 +192,7 @@ fun make_def sign fs (fname, ls, rec_name, tname) = let - val rhs = foldr (fn (T, t) => Abs ("", T, t)) + val rhs = Library.foldr (fn (T, t) => Abs ("", T, t)) ((map snd ls) @ [dummyT], list_comb (Const (rec_name, dummyT), fs @ map Bound (0 ::(length ls downto 1)))); @@ -216,10 +216,10 @@ let fun constrs_of (_, (_, _, cs)) = map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs; - val params_of = Library.assocs (flat (map constrs_of rec_eqns)); + val params_of = Library.assocs (List.concat (map constrs_of rec_eqns)); in induction - |> RuleCases.rename_params (map params_of (flat (map (map #1 o #3 o #2) descr))) + |> RuleCases.rename_params (map params_of (List.concat (map (map #1 o #3 o #2) descr))) |> RuleCases.save induction end; @@ -228,22 +228,22 @@ val (eqns, atts) = split_list eqns_atts; val sg = Theory.sign_of thy; val dt_info = DatatypePackage.get_datatypes thy; - val rec_eqns = foldr (process_eqn sg) (map snd eqns, []); + val rec_eqns = Library.foldr (process_eqn sg) (map snd eqns, []); val tnames = distinct (map (#1 o snd) rec_eqns); val dts = find_dts dt_info tnames tnames; val main_fns = map (fn (tname, {index, ...}) => (index, - fst (the (find_first (fn f => #1 (snd f) = tname) rec_eqns)))) + fst (valOf (find_first (fn f => #1 (snd f) = tname) rec_eqns)))) dts; val {descr, rec_names, rec_rewrites, ...} = if null dts then primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive") else snd (hd dts); - val (fnames, fnss) = foldr (process_fun sg descr rec_eqns) + val (fnames, fnss) = Library.foldr (process_fun sg descr rec_eqns) (main_fns, ([], [])); - val (fs, defs) = foldr (get_fns fnss) (descr ~~ rec_names, ([], [])); + val (fs, defs) = Library.foldr (get_fns fnss) (descr ~~ rec_names, ([], [])); val defs' = map (make_def sg fs) defs; val names1 = map snd fnames; val names2 = map fst rec_eqns; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/prop_logic.ML --- a/src/HOL/Tools/prop_logic.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/prop_logic.ML Thu Mar 03 12:43:01 2005 +0100 @@ -275,7 +275,7 @@ (* prop_formula list -> prop_formula *) - fun exists xs = foldl SOr (False, xs); + fun exists xs = Library.foldl SOr (False, xs); (* ------------------------------------------------------------------------- *) (* all: computes the conjunction over a list 'xs' of propositional formulas *) @@ -283,7 +283,7 @@ (* prop_formula list -> prop_formula *) - fun all xs = foldl SAnd (True, xs); + fun all xs = Library.foldl SAnd (True, xs); (* ------------------------------------------------------------------------- *) (* dot_product: ([x1,...,xn], [y1,...,yn]) -> x1*y1+...+xn*yn *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/recdef_package.ML --- a/src/HOL/Tools/recdef_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/recdef_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -91,7 +91,7 @@ val (del, rest) = Library.partition (Library.equal c o fst) congs; in if null del then (warning ("No recdef congruence rule for " ^ quote c); congs) else rest end; -val add_congs = curry (foldr (uncurry add_cong)); +val add_congs = curry (Library.foldr (uncurry add_cong)); end; @@ -255,7 +255,7 @@ val (thy, (simps' :: rules', [induct'])) = thy |> Theory.add_path bname - |> PureThy.add_thmss ((("simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) + |> PureThy.add_thmss ((("simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) |>>> PureThy.add_thms [(("induct", induct), [])]; val result = {simps = simps', rules = rules', induct = induct', tcs = tcs}; val thy = @@ -312,8 +312,8 @@ (case get_recdef thy name of NONE => error ("No recdef definition of constant: " ^ quote name) | SOME {tcs, ...} => tcs); - val i = if_none opt_i 1; - val tc = Library.nth_elem (i - 1, tcs) handle Library.LIST _ => + val i = getOpt (opt_i, 1); + val tc = List.nth (tcs, i - 1) handle Subscript => error ("No termination condition #" ^ string_of_int i ^ " in recdef definition of " ^ quote name); in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/recfun_codegen.ML --- a/src/HOL/Tools/recfun_codegen.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/recfun_codegen.ML Thu Mar 03 12:43:01 2005 +0100 @@ -44,7 +44,7 @@ in if Pattern.pattern (lhs_of thm) then (CodegenData.put (Symtab.update ((s, - thm :: if_none (Symtab.lookup (tab, s)) []), tab)) thy, thm) + thm :: getOpt (Symtab.lookup (tab, s), [])), tab)) thy, thm) else (warn thm; p) end handle TERM _ => (warn thm; p); @@ -70,7 +70,7 @@ (case Symtab.lookup (CodegenData.get thy, s) of NONE => [] | SOME thms => preprocess thy (del_redundant thy [] - (filter (fn thm => is_instance thy T + (List.filter (fn thm => is_instance thy T (snd (const_of (prop_of thm)))) thms))); fun mk_poly_id thy (s, T) = mk_const_id (sign_of thy) s ^ @@ -82,7 +82,7 @@ fun cycle g (xs, x) = if x mem xs then xs - else foldl (cycle g) (x :: xs, flat (Graph.find_paths g (x, x))); + else Library.foldl (cycle g) (x :: xs, List.concat (Graph.find_paths g (x, x))); fun add_rec_funs thy gr dep eqs = let @@ -125,10 +125,10 @@ if not (dep mem xs) then let val eqs'' = map (dest_eq o prop_of) - (flat (map (get_equations thy) cs)); + (List.concat (map (get_equations thy) cs)); val (gr3, fundef') = mk_fundef "" "fun " (Graph.add_edge (dname, dep) - (foldr (uncurry Graph.new_node) (map (fn k => + (Library.foldr (uncurry Graph.new_node) (map (fn k => (k, (SOME (EQN ("", dummyT, dname)), ""))) xs, Graph.del_nodes xs gr2))) eqs'' in put_code fundef' gr3 end diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/record_package.ML --- a/src/HOL/Tools/record_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/record_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -160,7 +160,7 @@ fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) = (case try (unsuffix ext_typeN) c_ext_type of NONE => raise TYPE ("RecordPackage.dest_recT", [typ], []) - | SOME c => ((c, Ts), last_elem Ts)) + | SOME c => ((c, Ts), List.last Ts)) | dest_recT typ = raise TYPE ("RecordPackage.dest_recT", [typ], []); fun is_recT T = @@ -180,8 +180,8 @@ fun rec_id i T = let val rTs = dest_recTs T - val rTs' = if i < 0 then rTs else take (i,rTs) - in foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end; + val rTs' = if i < 0 then rTs else Library.take (i,rTs) + in Library.foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end; (*** extend theory by record definition ***) @@ -426,11 +426,11 @@ val varifyT = map_type_tfree varify; val {records,extfields,...} = RecordsData.get_sg sg; val (flds,(more,_)) = split_last (Symtab.lookup_multi (extfields,name)); - val args = map varifyT (snd (#extension (the (Symtab.lookup (records,recname))))); + val args = map varifyT (snd (#extension (valOf (Symtab.lookup (records,recname))))); val tsig = Sign.tsig_of sg; fun unify (t,env) = Type.unify tsig env t; - val (subst,_) = foldr unify (but_last args ~~ but_last Ts,(Vartab.empty,0)); + val (subst,_) = Library.foldr unify (but_last args ~~ but_last Ts,(Vartab.empty,0)); val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds; in (flds',(more,moreT)) end; @@ -449,7 +449,7 @@ let val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} = RecordsData.get thy; - val fieldext' = foldl (fn (table,field) => Symtab.update_new ((field,extname_types),table)) + val fieldext' = Library.foldl (fn (table,field) => Symtab.update_new ((field,extname_types),table)) (fieldext,fields); val data=make_record_data records sel_upd equalities extinjects extsplit splits extfields fieldext'; @@ -473,11 +473,11 @@ fun bad_inst ((x, S), T) = if Sign.of_sort sign (T, S) then NONE else SOME x - val bads = mapfilter bad_inst (args ~~ types); + val bads = List.mapPartial bad_inst (args ~~ types); val inst = map fst args ~~ types; - val subst = Term.map_type_tfree (fn (x, _) => the (assoc (inst, x))); - val parent' = apsome (apfst (map subst)) parent; + val subst = Term.map_type_tfree (fn (x, _) => valOf (assoc (inst, x))); + val parent' = Option.map (apfst (map subst)) parent; val fields' = map (apsnd subst) fields; val extension' = apsnd (map subst) extension; in @@ -504,7 +504,7 @@ fun record_update_tr [t, u] = - foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) + Library.foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t) | record_update_tr ts = raise TERM ("record_update_tr", ts); fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts @@ -584,7 +584,7 @@ val (args,rest) = splitargs (map fst flds') fargs; val vartypes = map Type.varifyT types; val argtypes = map to_type args; - val (subst,_) = foldr unify (vartypes ~~ argtypes,(Vartab.empty,0)); + val (subst,_) = Library.foldr unify (vartypes ~~ argtypes,(Vartab.empty,0)); val alphas' = map ((Syntax.term_of_typ (! Syntax.show_sorts)) o (Envir.norm_type subst) o Type.varifyT) (but_last alphas); @@ -671,7 +671,7 @@ val flds' = Sign.extern sg Sign.constK f::map NameSpace.base fs; val (args',more) = split_last args; in (flds'~~args')@field_lst more end - handle LIST _ => [("",t)]) + handle UnequalLengths => [("",t)]) | NONE => [("",t)]) | NONE => [("",t)]) | _ => [("",t)]) @@ -777,12 +777,12 @@ ::map (apfst NameSpace.base) fs; val (args',more) = split_last args; val alphavars = map Type.varifyT (but_last alphas); - val (subst,_)= foldr unify (alphavars~~args',(Vartab.empty,0)); + val (subst,_)= Library.foldr unify (alphavars~~args',(Vartab.empty,0)); val flds'' =map (apsnd ((Envir.norm_type subst)o(Type.varifyT))) flds'; in flds''@field_lst more end - handle TUNIFY => [("",T)] - | LIST _=> [("",T)]) + handle TUNIFY => [("",T)] + | UnequalLengths => [("",T)]) | NONE => [("",T)]) | NONE => [("",T)]) | NONE => [("",T)]) @@ -840,7 +840,7 @@ let val {sel_upd={simpset,...},extsplit,...} = RecordsData.get_sg sg; val extsplits = - foldl (fn (thms,(n,_)) => (list (Symtab.lookup (extsplit,n)))@thms) + Library.foldl (fn (thms,(n,_)) => (list (Symtab.lookup (extsplit,n)))@thms) ([],dest_recTs T); val thms = (case get_splits sg (rec_id (~1) T) of SOME (all_thm,_,_,_) => @@ -854,7 +854,7 @@ local fun get_fields extfields T = - foldl (fn (xs,(eN,_))=>xs@(Symtab.lookup_multi (extfields,eN))) + Library.foldl (fn (xs,(eN,_))=>xs@(Symtab.lookup_multi (extfields,eN))) ([],(dest_recTs T)); in (* record_simproc *) @@ -974,10 +974,10 @@ * updates again, the optimised term is constructed. *) fun mk_updterm upds already (t as ((upd as Const (u,uT)) $ k $ r)) = - if is_some (Symtab.lookup (upds,u)) + if isSome (Symtab.lookup (upds,u)) then let fun rest already = mk_updterm upds already - in if is_some (Symtab.lookup (already,u)) + in if isSome (Symtab.lookup (already,u)) then (case (rest already r) of Init ((sprout,skel),vars) => let @@ -1146,8 +1146,8 @@ (s = "all" orelse s = "All" orelse s = "Ex") andalso is_recT T | _ => false); - val goal = Library.nth_elem (i - 1, Thm.prems_of st); - val frees = filter (is_recT o type_of) (term_frees goal); + val goal = List.nth (Thm.prems_of st, i - 1); + val frees = List.filter (is_recT o type_of) (term_frees goal); fun mk_split_free_tac free induct_thm i = let val cfree = cterm_of sg free; @@ -1172,13 +1172,13 @@ end) | split_free_tac _ _ _ = NONE; - val split_frees_tacs = mapfilter (split_free_tac P i) frees; + val split_frees_tacs = List.mapPartial (split_free_tac P i) frees; val simprocs = if has_rec goal then [record_split_simproc P] else []; in st |> (EVERY split_frees_tacs) THEN (Simplifier.full_simp_tac (simpset addsimps thms addsimprocs simprocs) i) - end handle Library.LIST _ => Seq.empty; + end handle Empty => Seq.empty; end; @@ -1193,7 +1193,7 @@ (s = "all" orelse s = "All") andalso is_recT T | _ => false); - val goal = Library.nth_elem (i - 1, Thm.prems_of st); + val goal = List.nth (Thm.prems_of st, i - 1); fun is_all t = (case t of (Const (quantifier, _)$_) => @@ -1204,7 +1204,7 @@ then Simplifier.full_simp_tac (HOL_basic_ss addsimprocs [record_split_simproc is_all]) i st else Seq.empty - end handle Library.LIST _ => Seq.empty; + end handle Subscript => Seq.empty; (* wrapper *) @@ -1251,14 +1251,14 @@ fun try_param_tac s rule i st = let val cert = cterm_of (Thm.sign_of_thm st); - val g = nth_elem (i - 1, prems_of st); + val g = List.nth (prems_of st, i - 1); val params = Logic.strip_params g; val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g); val rule' = Thm.lift_rule (st, i) rule; val (P, ys) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_assums_concl (prop_of rule'))); (* ca indicates if rule is a case analysis or induction rule *) - val (x, ca) = (case rev (drop (length params, ys)) of + val (x, ca) = (case rev (Library.drop (length params, ys)) of [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true) | [x] => (head_of x, false)); @@ -1278,13 +1278,13 @@ fun ex_inst_tac i st = let val sg = sign_of_thm st; - val g = nth_elem (i - 1, prems_of st); + val g = List.nth (prems_of st, i - 1); val params = Logic.strip_params g; val exI' = Thm.lift_rule (st, i) exI; val (_$(_$x)) = Logic.strip_assums_concl (hd (prems_of exI')); val cx = cterm_of sg (fst (strip_comb x)); - in Seq.single (foldl (fn (st,v) => + in Seq.single (Library.foldl (fn (st,v) => Seq.hd (compose_tac (false, cterm_instantiate [(cx,cterm_of sg (list_abs (params,Bound v)))] exI',1) @@ -1307,7 +1307,7 @@ fun mixit convs refls = let fun f ((res,lhs,rhs),refl) = ((refl,List.revAppend (lhs,refl::tl rhs))::res,hd rhs::lhs,tl rhs); - in #1 (foldl f (([],[],convs),refls)) end; + in #1 (Library.foldl f (([],[],convs),refls)) end; fun extension_definition full name fields names alphas zeta moreT more vars thy = let @@ -1332,7 +1332,7 @@ val ext_decl = (mk_extC (name,extT) fields_moreTs); (* val ext_spec = Const ext_decl :== - (foldr (uncurry lambda) + (Library.foldr (uncurry lambda) (vars@[more],(mk_Abs name repT extT $ (foldr1 HOLogic.mk_prod (vars@[more]))))) *) val ext_spec = list_comb (Const ext_decl,vars@[more]) :== @@ -1484,7 +1484,7 @@ (* freezeT just for performance, to adjust the maxidx, so that nodup_vars will not be called during combination *) fun mkthm (udef,(fld_refl,thms)) = - let val bdyeq = foldl (uncurry Thm.combination) (constr_refl,thms); + let val bdyeq = Library.foldl (uncurry Thm.combination) (constr_refl,thms); (* (|N=N (|N=N,M=M,K=K,more=more|) M=M (|N=N,M=M,K=K,more=more|) K=K' @@ -1537,7 +1537,7 @@ fun chunks [] [] = [] | chunks [] xs = [xs] - | chunks (l::ls) xs = take (l,xs)::chunks ls (drop (l,xs)); + | chunks (l::ls) xs = Library.take (l,xs)::chunks ls (Library.drop (l,xs)); fun chop_last [] = error "last: list should not be empty" | chop_last [x] = ([],x) @@ -1551,7 +1551,7 @@ * of parent extensions, starting with the root of the record hierarchy *) fun mk_recordT extT parent_exts = - foldr (fn ((parent,Ts),T) => Type (parent, subst_last T Ts)) (parent_exts,extT); + Library.foldr (fn ((parent,Ts),T) => Type (parent, subst_last T Ts)) (parent_exts,extT); @@ -1591,7 +1591,7 @@ val (bfields, field_syntax) = split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields); - val parent_fields = flat (map #fields parents); + val parent_fields = List.concat (map #fields parents); val parent_chunks = map (length o #fields) parents; val parent_names = map fst parent_fields; val parent_types = map snd parent_fields; @@ -1605,7 +1605,7 @@ val names = map fst fields; val extN = full bname; val types = map snd fields; - val alphas_fields = foldr add_typ_tfree_names (types,[]); + val alphas_fields = Library.foldr add_typ_tfree_names (types,[]); val alphas_ext = alphas inter alphas_fields; val len = length fields; val variants = variantlist (map fst bfields, moreN::rN::rN ^ "'"::wN::parent_variants); @@ -1647,7 +1647,7 @@ val extension = let val (n,Ts) = extension_scheme in (n,subst_last HOLogic.unitT Ts) end; val extension_names = (map ((unsuffix ext_typeN) o fst o #extension) parents) @ [extN]; - val extension_id = foldl (op ^) ("",extension_names); + val extension_id = Library.foldl (op ^) ("",extension_names); fun rec_schemeT n = mk_recordT extT (map #extension (prune n parents)); @@ -1663,7 +1663,7 @@ let val (args',more) = chop_last args; fun mk_ext' (((name,T),args),more) = mk_ext (name,T) (args@[more]); fun build Ts = - foldr mk_ext' (prune n (extension_names ~~ Ts ~~ (chunks parent_chunks args')),more) + Library.foldr mk_ext' (prune n (extension_names ~~ Ts ~~ (chunks parent_chunks args')),more) in if more = HOLogic.unit then build (map recT (0 upto parent_len)) @@ -1681,7 +1681,7 @@ (* prepare print translation functions *) val field_tr's = - print_translation (distinct (flat (map NameSpace.accesses' (full_moreN :: names)))); + print_translation (distinct (List.concat (map NameSpace.accesses' (full_moreN :: names)))); val adv_ext_tr's = let @@ -1822,13 +1822,13 @@ end; val split_object_prop = - let fun ALL vs t = foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) (vs,t) + let fun ALL vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) (vs,t) in (ALL [dest_Free r0] (P $ r0)) === (ALL (map dest_Free all_vars_more) (P $ r_rec0)) end; val split_ex_prop = - let fun EX vs t = foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) (vs,t) + let fun EX vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) (vs,t) in (EX [dest_Free r0] (P $ r0)) === (EX (map dest_Free all_vars_more) (P $ r_rec0)) end; @@ -2041,14 +2041,14 @@ fun prep_inst T = snd (cert_typ sign ([], T)); - val parent = apsome (apfst (map prep_inst) o prep_raw_parent sign) raw_parent + val parent = Option.map (apfst (map prep_inst) o prep_raw_parent sign) raw_parent handle ERROR => error ("The error(s) above in parent record specification"); val parents = add_parents thy parent []; val init_env = (case parent of NONE => [] - | SOME (types, _) => foldr Term.add_typ_tfrees (types, [])); + | SOME (types, _) => Library.foldr Term.add_typ_tfrees (types, [])); (* fields *) @@ -2065,7 +2065,7 @@ (* args *) val defaultS = Sign.defaultS sign; - val args = map (fn x => (x, if_none (assoc (envir, x)) defaultS)) params; + val args = map (fn x => (x, getOpt (assoc (envir, x), defaultS))) params; (* errors *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/refute.ML --- a/src/HOL/Tools/refute.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/refute.ML Thu Mar 03 12:43:01 2005 +0100 @@ -93,7 +93,7 @@ fun tree_foldl f = let fun itl (e, Leaf x) = f(e,x) - | itl (e, Node xs) = foldl (tree_foldl f) (e,xs) + | itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs) in itl end; @@ -192,7 +192,7 @@ parameters = Symtab.merge (op=) (pa1, pa2)}; fun print sg {interpreters, printers, parameters} = Pretty.writeln (Pretty.chunks - [Pretty.strs ("default parameters:" :: flat (map (fn (name,value) => [name, "=", value]) (Symtab.dest parameters))), + [Pretty.strs ("default parameters:" :: List.concat (map (fn (name,value) => [name, "=", value]) (Symtab.dest parameters))), Pretty.strs ("interpreters:" :: map fst interpreters), Pretty.strs ("printers:" :: map fst printers)]); end; @@ -250,7 +250,7 @@ if null terms then "empty interpretation (no free variables in term)\n" else - space_implode "\n" (mapfilter (fn (t,intr) => + space_implode "\n" (List.mapPartial (fn (t,intr) => (* print constants only if 'show_consts' is true *) if (!show_consts) orelse not (is_Const t) then SOME (Sign.string_of_term (sign_of thy) t ^ ": " ^ Sign.string_of_term (sign_of thy) (print thy model t intr assignment)) @@ -359,9 +359,9 @@ (* TODO: it is currently not possible to specify a size for a type *) (* whose name is one of the other parameters (e.g. 'maxvars') *) (* (string * int) list *) - val sizes = mapfilter + val sizes = List.mapPartial (fn (name,value) => (case Int.fromString value of SOME i => SOME (name, i) | NONE => NONE)) - (filter (fn (name,_) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" andalso name<>"satsolver") + (List.filter (fn (name,_) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" andalso name<>"satsolver") allparams) in {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars, maxtime=maxtime, satsolver=satsolver} @@ -382,12 +382,12 @@ fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) = (* replace a 'DtTFree' variable by the associated type *) - (the o assoc) (typ_assoc, DatatypeAux.DtTFree a) + (valOf o assoc) (typ_assoc, DatatypeAux.DtTFree a) | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) = Type (s, map (typ_of_dtyp descr typ_assoc) ds) | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) = let - val (s, ds, _) = (the o assoc) (descr, i) + val (s, ds, _) = (valOf o assoc) (descr, i) in Type (s, map (typ_of_dtyp descr typ_assoc) ds) end; @@ -419,7 +419,7 @@ let val _ = immediate_output "Adding axioms..." (* (string * Term.term) list *) - val axioms = flat (map (Symtab.dest o #axioms o Theory.rep_theory) (thy :: Theory.ancestors_of thy)) + val axioms = List.concat (map (Symtab.dest o #axioms o Theory.rep_theory) (thy :: Theory.ancestors_of thy)) (* string list *) val rec_names = Symtab.foldl (fn (acc, (_, info)) => #rec_names info @ acc) ([], DatatypePackage.get_datatypes thy) @@ -478,7 +478,7 @@ let (* obtain the axioms generated by the "axclass" command *) (* (string * Term.term) list *) - val class_axioms = filter (fn (s, _) => String.isPrefix (class ^ ".axioms_") s) axioms + val class_axioms = List.filter (fn (s, _) => String.isPrefix (class ^ ".axioms_") s) axioms (* replace the one schematic type variable in each axiom by the actual type 'T' *) (* (string * Term.term) list *) val monomorphic_class_axioms = map (fn (axname, ax) => @@ -498,7 +498,7 @@ collect_term_axioms (ax :: axs, ax) ) in - foldl collect_axiom (axs, monomorphic_class_axioms) + Library.foldl collect_axiom (axs, monomorphic_class_axioms) end (* string list *) val sort = (case T of @@ -509,7 +509,7 @@ (* string list *) val superclasses = Graph.all_succs ((#classes o Type.rep_tsig o Sign.tsig_of o sign_of) thy) sort in - foldl collect_class_axioms (axs, superclasses) + Library.foldl collect_class_axioms (axs, superclasses) end (* Term.term list * Term.typ -> Term.term list *) and collect_type_axioms (axs, T) = @@ -557,20 +557,20 @@ case DatatypePackage.datatype_info thy s of SOME info => (* inductive datatype *) (* only collect relevant type axioms for the argument types *) - foldl collect_type_axioms (axs, Ts) + Library.foldl collect_type_axioms (axs, Ts) | NONE => (case get_typedefn axioms of SOME (axname, ax) => if mem_term (ax, axs) then (* only collect relevant type axioms for the argument types *) - foldl collect_type_axioms (axs, Ts) + Library.foldl collect_type_axioms (axs, Ts) else (immediate_output (" " ^ axname); collect_term_axioms (ax :: axs, ax)) | NONE => (* unspecified type, perhaps introduced with 'typedecl' *) (* at least collect relevant type axioms for the argument types *) - foldl collect_type_axioms (axs, Ts)) + Library.foldl collect_type_axioms (axs, Ts)) end | TFree _ => collect_sort_axioms (axs, T) (* axiomatic type classes *) | TVar _ => collect_sort_axioms (axs, T) (* axiomatic type classes *) @@ -590,7 +590,7 @@ | Const ("arbitrary", T) => collect_type_axioms (axs, T) | Const ("The", T) => let - val ax = specialize_type (("The", T), (the o assoc) (axioms, "HOL.the_eq_trivial")) + val ax = specialize_type (("The", T), (valOf o assoc) (axioms, "HOL.the_eq_trivial")) in if mem_term (ax, axs) then collect_type_axioms (axs, T) @@ -600,7 +600,7 @@ end | Const ("Hilbert_Choice.Eps", T) => let - val ax = specialize_type (("Hilbert_Choice.Eps", T), (the o assoc) (axioms, "Hilbert_Choice.someI")) + val ax = specialize_type (("Hilbert_Choice.Eps", T), (valOf o assoc) (axioms, "Hilbert_Choice.someI")) in if mem_term (ax, axs) then collect_type_axioms (axs, T) @@ -686,7 +686,7 @@ val inclass = Logic.mk_inclass (TVar (("'a", 0), [class]), class) (* Term.term option *) val ofclass_ax = (SOME (specialize_type ((s, T), inclass)) handle Type.TYPE_MATCH => NONE) - val intro_ax = (apsome (fn t => specialize_type ((s, T), t)) (assoc (axioms, class ^ ".intro")) handle Type.TYPE_MATCH => NONE) + val intro_ax = (Option.map (fn t => specialize_type ((s, T), t)) (assoc (axioms, class ^ ".intro")) handle Type.TYPE_MATCH => NONE) val axs' = (case ofclass_ax of NONE => axs | SOME ax => if mem_term (ax, axs) then (* collect relevant type axioms *) collect_type_axioms (axs, T) @@ -734,7 +734,7 @@ (* (Term.indexname * Term.typ) list *) val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t)) in - foldl + Library.foldl (fn (t', ((x, i), T)) => (Term.all T) $ Abs (x, T, abstract_over (Var((x, i), T), t'))) (t, vars) end @@ -772,7 +772,7 @@ let val index = #index info val descr = #descr info - val (_, dtyps, constrs) = (the o assoc) (descr, index) + val (_, dtyps, constrs) = (valOf o assoc) (descr, index) val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = (if Library.exists (fn d => @@ -787,14 +787,14 @@ else acc) (* collect argument types *) - val acc_args = foldr collect_types (Ts, acc') + val acc_args = Library.foldr collect_types (Ts, acc') (* collect constructor types *) - val acc_constrs = foldr collect_types (flat (map (fn (_, ds) => map (typ_of_dtyp descr typ_assoc) ds) constrs), acc_args) + val acc_constrs = Library.foldr collect_types (List.concat (map (fn (_, ds) => map (typ_of_dtyp descr typ_assoc) ds) constrs), acc_args) in acc_constrs end | NONE => (* not an inductive datatype, e.g. defined via "typedef" or "typedecl" *) - T ins (foldr collect_types (Ts, acc))) + T ins (Library.foldr collect_types (Ts, acc))) | TFree _ => T ins acc | TVar _ => T ins acc) in @@ -853,9 +853,9 @@ NONE | make_first max len sum = if sum<=max orelse max<0 then - apsome (fn xs' => sum :: xs') (make_first max (len-1) 0) + Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0) else - apsome (fn xs' => max :: xs') (make_first max (len-1) (sum-max)) + Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max)) (* enumerates all int lists with a fixed length, where 0<=x<='max' for *) (* all list elements x (unless 'max'<0) *) (* int -> int -> int -> int list -> int list option *) @@ -867,12 +867,12 @@ | next max len sum (x1::x2::xs) = if x1>0 andalso (x2 assoc (sizes, string_of_typ T) = NONE) xs + val mutables = List.filter (fn (T, _) => assoc (sizes, string_of_typ T) = NONE) xs (* subtract 'minsize' from every size (will be added again at the end) *) val diffs = map (fn (_, n) => n-minsize) mutables in @@ -931,7 +931,7 @@ (* Term.term list *) val axioms = collect_axioms thy t (* Term.typ list *) - val types = foldl (fn (acc, t') => acc union (ground_types thy t')) ([], t :: axioms) + val types = Library.foldl (fn (acc, t') => acc union (ground_types thy t')) ([], t :: axioms) val _ = writeln ("Ground types: " ^ (if null types then "none." else commas (map (Sign.string_of_typ (sign_of thy)) types))) @@ -945,7 +945,7 @@ let val index = #index info val descr = #descr info - val (_, _, constrs) = (the o assoc) (descr, index) + val (_, _, constrs) = (valOf o assoc) (descr, index) in (* recursive datatype? *) Library.exists (fn (_, ds) => Library.exists DatatypeAux.is_rec_type ds) constrs @@ -1053,7 +1053,7 @@ (* (Term.indexname * Term.typ) list *) val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t)) (* Term.term *) - val ex_closure = foldl + val ex_closure = Library.foldl (fn (t', ((x,i),T)) => (HOLogic.exists_const T) $ Abs (x, T, abstract_over (Var((x,i),T), t'))) (t, vars) (* If 't' is of type 'propT' (rather than 'boolT'), applying *) @@ -1075,7 +1075,7 @@ (* theory -> (string * string) list -> Thm.thm -> int -> unit *) fun refute_subgoal thy params thm subgoal = - refute_term thy params (nth_elem (subgoal, prems_of thm)); + refute_term thy params (List.nth (prems_of thm, subgoal)); (* ------------------------------------------------------------------------- *) @@ -1116,7 +1116,7 @@ map (fn x => [x]) xs | pick_all n xs = let val rec_pick = pick_all (n-1) xs in - foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs) + Library.foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs) end in case intr of @@ -1269,7 +1269,7 @@ map (fn x => [x]) xs | pick_all (xs::xss) = let val rec_pick = pick_all xss in - foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs) + Library.foldl (fn (acc,x) => (cons_list x rec_pick) @ acc) ([],xs) end | pick_all _ = raise REFUTE ("interpretation_apply", "empty list (in pick_all)") @@ -1296,8 +1296,8 @@ let val Ts = binder_types (fastype_of t) in - foldr (fn (T, t) => Abs ("", T, t)) - (take (i, Ts), list_comb (t, map Bound (i-1 downto 0))) + Library.foldr (fn (T, t) => Abs ("", T, t)) + (Library.take (i, Ts), list_comb (t, map Bound (i-1 downto 0))) end; (* ------------------------------------------------------------------------- *) @@ -1306,7 +1306,7 @@ (* int list -> int *) - fun sum xs = foldl op+ (0, xs); + fun sum xs = Library.foldl op+ (0, xs); (* ------------------------------------------------------------------------- *) (* product: returns the product of a list 'xs' of integers *) @@ -1314,7 +1314,7 @@ (* int list -> int *) - fun product xs = foldl op* (1, xs); + fun product xs = Library.foldl op* (1, xs); (* ------------------------------------------------------------------------- *) (* size_of_dtyp: the size of (an initial fragment of) an inductive data type *) @@ -1354,7 +1354,7 @@ (* unit -> (interpretation * model * arguments) option *) fun interpret_groundtype () = let - val size = (if T = Term.propT then 2 else (the o assoc) (typs, T)) (* the model MUST specify a size for ground types *) + val size = (if T = Term.propT then 2 else (valOf o assoc) (typs, T)) (* the model MUST specify a size for ground types *) val next = next_idx+size val _ = (if next-1>maxvars andalso maxvars>0 then raise MAXVARS_EXCEEDED else ()) (* check if 'maxvars' is large enough *) (* prop_formula list *) @@ -1416,7 +1416,7 @@ | Var (_, T) => interpret_groundterm T | Bound i => - SOME (nth_elem (i, #bounds args), model, args) + SOME (List.nth (#bounds args, i), model, args) | Abs (x, T, body) => let (* create all constants of type 'T' *) @@ -1656,7 +1656,7 @@ let val index = #index info val descr = #descr info - val (_, dtyps, constrs) = (the o assoc) (descr, index) + val (_, dtyps, constrs) = (valOf o assoc) (descr, index) val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = (if Library.exists (fn d => @@ -1723,7 +1723,7 @@ let val index = #index info val descr = #descr info - val (_, dtyps, constrs) = (the o assoc) (descr, index) + val (_, dtyps, constrs) = (valOf o assoc) (descr, index) val typ_assoc = dtyps ~~ Ts' (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = (if Library.exists (fn d => @@ -1846,7 +1846,7 @@ (* after it generate further elements *) val offset = size_of_dtyp thy typs' descr typ_assoc constrs1 + (if is_rec_constr ctypes then - size_of_dtyp thy typs' descr typ_assoc (filter (not o is_rec_constr o snd) cs) + size_of_dtyp thy typs' descr typ_assoc (List.filter (not o is_rec_constr o snd) cs) else 0) in @@ -1904,7 +1904,7 @@ let val index = #index info val descr = #descr info - val (_, dtyps, constrs) = (the o assoc) (descr, index) + val (_, dtyps, constrs) = (valOf o assoc) (descr, index) (* the total number of constructors, including those of different *) (* (mutually recursive) datatypes within the same descriptor 'descr' *) val constrs_count = sum (map (fn (_, (_, _, cs)) => length cs) descr) @@ -1931,7 +1931,7 @@ end) ((model, args), params) val (typs, terms) = model' (* the type of a recursion operator: [T1, ..., Tn, IDT] ---> Tresult *) - val IDT = nth_elem (constrs_count, binder_types T) + val IDT = List.nth (binder_types T, constrs_count) val typ_assoc = dtyps ~~ (snd o dest_Type) IDT (* interpret each constructor of the datatype *) (* TODO: we probably need to interpret every constructor in the descriptor, *) @@ -1993,13 +1993,13 @@ | select_subtree (Leaf _, _) = raise REFUTE ("IDT_recursion_interpreter", "interpretation for parameter is a leaf; cannot select a subtree") | select_subtree (Node tr, x::xs) = - select_subtree (nth_elem (x, tr), xs) + select_subtree (List.nth (tr, x), xs) (* select the correct subtree of the parameter corresponding to constructor 'c' *) - val p_intr = select_subtree (nth_elem (c, p_intrs), args) + val p_intr = select_subtree (List.nth (p_intrs, c), args) (* find the indices of recursive arguments *) - val rec_args = map snd (filter (DatatypeAux.is_rec_type o fst) ((snd (nth_elem (c, constrs))) ~~ args)) + val rec_args = map snd (List.filter (DatatypeAux.is_rec_type o fst) ((snd (List.nth (constrs, c))) ~~ args)) (* apply 'p_intr' to recursively computed results *) - val rec_p_intr = foldl (fn (i, n) => interpretation_apply (i, Array.sub (INTRS, n))) (p_intr, rec_args) + val rec_p_intr = Library.foldl (fn (i, n) => interpretation_apply (i, Array.sub (INTRS, n))) (p_intr, rec_args) (* update 'INTRS' *) val _ = Array.update (INTRS, elem, rec_p_intr) in @@ -2036,7 +2036,7 @@ val constants = make_constants i_set (* interpretation -> int *) fun number_of_elements (Node xs) = - foldl (fn (n, x) => + Library.foldl (fn (n, x) => if x=TT then n+1 else if x=FF then @@ -2214,7 +2214,7 @@ val HOLogic_empty_set = Const ("{}", HOLogic_setT) val HOLogic_insert = Const ("insert", HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) in - SOME (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) (pairs, HOLogic_empty_set)) + SOME (Library.foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) (pairs, HOLogic_empty_set)) end | Type ("prop", []) => (case index_from_interpretation intr of @@ -2259,7 +2259,7 @@ Node xs => xs | _ => raise REFUTE ("set_printer", "interpretation for set type is a leaf")) (* Term.term list *) - val elements = mapfilter (fn (arg, result) => + val elements = List.mapPartial (fn (arg, result) => case result of Leaf [fmTrue, fmFalse] => if PropLogic.eval assignment fmTrue then @@ -2275,7 +2275,7 @@ val HOLogic_empty_set = Const ("{}", HOLogic_setT) val HOLogic_insert = Const ("insert", T --> HOLogic_setT --> HOLogic_setT) in - SOME (foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc) (HOLogic_empty_set, elements)) + SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc) (HOLogic_empty_set, elements)) end | _ => NONE @@ -2299,7 +2299,7 @@ val (typs, _) = model val index = #index info val descr = #descr info - val (_, dtyps, constrs) = (the o assoc) (descr, index) + val (_, dtyps, constrs) = (valOf o assoc) (descr, index) val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = (if Library.exists (fn d => @@ -2345,7 +2345,7 @@ search (xs, 0) end in - apsome (fn args => (cTerm, cargs, args)) (get_args iC) + Option.map (fn args => (cTerm, cargs, args)) (get_args iC) end (* Term.term * DatatypeAux.dtyp list * int list *) val (cTerm, cargs, args) = (case get_first get_constr_args constrs of @@ -2359,10 +2359,10 @@ (* list, so there might be a more efficient implementation that does *) (* not generate all constants *) in - print thy (typs, []) (Free ("dummy", dT)) (nth_elem (n, consts)) assignment + print thy (typs, []) (Free ("dummy", dT)) (List.nth (consts, n)) assignment end) (cargs ~~ args) in - SOME (foldl op$ (cTerm, argsTerms)) + SOME (Library.foldl op$ (cTerm, argsTerms)) end end | NONE => (* not an inductive datatype *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/refute_isar.ML --- a/src/HOL/Tools/refute_isar.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/refute_isar.ML Thu Mar 03 12:43:01 2005 +0100 @@ -48,7 +48,7 @@ val thy = (Toplevel.theory_of state) val thm = ((snd o snd) (Proof.get_goal (Toplevel.proof_of state))) in - Refute.refute_subgoal thy (if_none parms []) thm ((if_none subgoal 1)-1) + Refute.refute_subgoal thy (getOpt (parms, [])) thm ((getOpt (subgoal, 1))-1) end) ); @@ -79,7 +79,7 @@ in Toplevel.theory (fn thy => let - val thy' = add_params (thy, (if_none args [])) + val thy' = add_params (thy, (getOpt (args, []))) val new_default_params = Refute.get_default_params thy' val output = if new_default_params=[] then "none" else (space_implode "\n" (map (fn (name,value) => name ^ "=" ^ value) new_default_params)) in diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/rewrite_hol_proof.ML --- a/src/HOL/Tools/rewrite_hol_proof.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/rewrite_hol_proof.ML Thu Mar 03 12:43:01 2005 +0100 @@ -311,15 +311,15 @@ fun mk_AbsP P t = AbsP ("H", P, t); fun elim_cong Ts (PThm (("HOL.iffD1", _), _, _, _) % _ % _ %% prf1 %% prf2) = - apsome (make_subst Ts prf2 []) (strip_cong [] prf1) + Option.map (make_subst Ts prf2 []) (strip_cong [] prf1) | elim_cong Ts (PThm (("HOL.iffD1", _), _, _, _) % P % _ %% prf) = - apsome (mk_AbsP P o make_subst Ts (PBound 0) []) + Option.map (mk_AbsP P o make_subst Ts (PBound 0) []) (strip_cong [] (incr_pboundvars 1 0 prf)) | elim_cong Ts (PThm (("HOL.iffD2", _), _, _, _) % _ % _ %% prf1 %% prf2) = - apsome (make_subst Ts prf2 [] o + Option.map (make_subst Ts prf2 [] o apsnd (map (make_sym Ts))) (strip_cong [] prf1) | elim_cong Ts (PThm (("HOL.iffD2", _), _, _, _) % _ % P %% prf) = - apsome (mk_AbsP P o make_subst Ts (PBound 0) [] o + Option.map (mk_AbsP P o make_subst Ts (PBound 0) [] o apsnd (map (make_sym Ts))) (strip_cong [] (incr_pboundvars 1 0 prf)) | elim_cong _ _ = NONE; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/sat_solver.ML --- a/src/HOL/Tools/sat_solver.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/sat_solver.ML Thu Mar 03 12:43:01 2005 +0100 @@ -193,7 +193,7 @@ NONE (* string -> int list *) fun int_list_from_string s = - mapfilter (option o Int.fromString) (space_explode " " s) + List.mapPartial (option o Int.fromString) (space_explode " " s) (* int list -> assignment *) fun assignment_from_list [] i = NONE (* the SAT solver didn't provide a value for this variable *) @@ -232,7 +232,7 @@ else parse_lines lines in - (parse_lines o (filter (fn l => l <> "")) o split_lines o File.read) path + (parse_lines o (List.filter (fn l => l <> "")) o split_lines o File.read) path end; (* ------------------------------------------------------------------------- *) @@ -310,10 +310,10 @@ o (map (map literal_from_int)) o clauses o (map int_from_string) - o flat + o List.concat o (map (String.fields (fn c => c mem [#" ", #"\t", #"\n"]))) o filter_preamble - o (filter (fn l => l <> "")) + o (List.filter (fn l => l <> "")) o split_lines o File.read) path @@ -343,7 +343,7 @@ (* string -> solver *) fun invoke_solver name = - (the o assoc) (!solvers, name); + (valOf o assoc) (!solvers, name); end; (* SatSolver *) @@ -458,7 +458,7 @@ | False => NONE | _ => let - val x = the (fresh_var xs') (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *) + val x = valOf (fresh_var xs') (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *) in case dpll (x::xs') fm' of (* passing fm' rather than fm should save some simplification work *) SOME xs'' => SOME xs'' diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/specification_package.ML --- a/src/HOL/Tools/specification_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/specification_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -141,7 +141,7 @@ val tsig = Sign.tsig_of (sign_of thy) val _ = assert (forall (fn v => Type.of_sort tsig (type_of v,HOLogic.typeS)) frees) "Specificaton: Only free variables of sort 'type' allowed" - val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) (map dest_Free frees,prop) + val prop_closed = Library.foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) (map dest_Free frees,prop) in (prop_closed,frees) end @@ -164,7 +164,7 @@ val c' = fst (Type.varify (c,[])) val (cname,ctyp) = dest_Const c' in - case filter (fn t => let val (name,typ) = dest_Const t + case List.filter (fn t => let val (name,typ) = dest_Const t in name = cname andalso typ_equiv typ ctyp end) thawed_prop_consts of [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term sg c) ^ "\" found") @@ -182,7 +182,7 @@ in HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop)) end - val ex_prop = foldr mk_exist (proc_consts,prop) + val ex_prop = Library.foldr mk_exist (proc_consts,prop) val cnames = map (fst o dest_Const) proc_consts fun post_process (arg as (thy,thm)) = let @@ -194,7 +194,7 @@ in thm RS spec' end - fun remove_alls frees thm = foldl (inst_all (sign_of_thm thm)) (thm,frees) + fun remove_alls frees thm = Library.foldl (inst_all (sign_of_thm thm)) (thm,frees) fun process_single ((name,atts),rew_imp,frees) args = let fun undo_imps thm = diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/split_rule.ML Thu Mar 03 12:43:01 2005 +0100 @@ -66,7 +66,7 @@ (* complete splitting of partially splitted rules *) fun ap_split' (T::Ts) U u = Abs ("v", T, ap_split' Ts U - (ap_split T (flat (map HOLogic.prodT_factors Ts) ---> U) + (ap_split T (List.concat (map HOLogic.prodT_factors Ts) ---> U) (incr_boundvars 1 u) $ Bound 0)) | ap_split' _ _ u = u; @@ -74,9 +74,9 @@ let val cterm = Thm.cterm_of (#sign (Thm.rep_thm rl)) val (Us', U') = strip_type T; - val Us = take (length ts, Us'); - val U = drop (length ts, Us') ---> U'; - val T' = flat (map HOLogic.prodT_factors Us) ---> U; + val Us = Library.take (length ts, Us'); + val U = Library.drop (length ts, Us') ---> U'; + val T' = List.concat (map HOLogic.prodT_factors Us) ---> U; fun mk_tuple ((xs, insts), v as Var ((a, _), T)) = let val Ts = HOLogic.prodT_factors T; @@ -87,7 +87,7 @@ | mk_tuple (x, _) = x; val newt = ap_split' Us U (Var (v, T')); val cterm = Thm.cterm_of (#sign (Thm.rep_thm rl)); - val (vs', insts) = foldl mk_tuple ((vs, []), ts); + val (vs', insts) = Library.foldl mk_tuple ((vs, []), ts); in (instantiate ([], [(cterm t, cterm newt)] @ insts) rl, vs') end @@ -96,19 +96,19 @@ fun collect_vars (vs, Abs (_, _, t)) = collect_vars (vs, t) | collect_vars (vs, t) = (case strip_comb t of (v as Var _, ts) => (v, ts)::vs - | (t, ts) => foldl collect_vars (vs, ts)); + | (t, ts) => Library.foldl collect_vars (vs, ts)); val split_rule_var = Drule.standard o remove_internal_split o split_rule_var'; (*curries ALL function variables occurring in a rule's conclusion*) fun split_rule rl = - foldr split_rule_var' (Term.term_vars (concl_of rl), rl) + Library.foldr split_rule_var' (Term.term_vars (concl_of rl), rl) |> remove_internal_split |> Drule.standard; fun complete_split_rule rl = - fst (foldr complete_split_rule_var + fst (Library.foldr complete_split_rule_var (collect_vars ([], concl_of rl), (rl, map (fst o fst o dest_Var) (Term.term_vars (#prop (Thm.rep_thm rl)))))) |> remove_internal_split @@ -121,7 +121,7 @@ fun split_rule_goal xss rl = let fun one_split i (th, s) = Tactic.rule_by_tactic (pair_tac s i) th; - fun one_goal (xs, i) th = foldl (one_split i) (th, xs); + fun one_goal (xs, i) th = Library.foldl (one_split i) (th, xs); in Drule.standard (Simplifier.full_simplify split_rule_ss (Library.foldln one_goal xss rl)) |> RuleCases.save rl diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/Tools/typedef_package.ML --- a/src/HOL/Tools/typedef_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/Tools/typedef_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -104,9 +104,9 @@ val thms = PureThy.get_thmss thy (map (rpair NONE) witn_names) @ witn_thms; val tac = witn1_tac THEN - TRY (rewrite_goals_tac (filter is_def thms)) THEN + TRY (rewrite_goals_tac (List.filter is_def thms)) THEN TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN - if_none witn2_tac (TRY (ALLGOALS (CLASET' blast_tac))); + getOpt (witn2_tac, TRY (ALLGOALS (CLASET' blast_tac))); in message ("Proving non-emptiness of set " ^ quote (string_of_cterm cset) ^ " ..."); Tactic.prove (Theory.sign_of thy) [] [] goal (K tac) @@ -144,12 +144,12 @@ val goal_pat = mk_nonempty (Var (vname, setT)); (*lhs*) - val lhs_tfrees = map (fn v => (v, if_none (assoc (rhs_tfrees, v)) HOLogic.typeS)) vs; + val lhs_tfrees = map (fn v => (v, getOpt (assoc (rhs_tfrees, v), HOLogic.typeS))) vs; val tname = Syntax.type_name t mx; val full_tname = full tname; val newT = Type (full_tname, map TFree lhs_tfrees); - val (Rep_name, Abs_name) = if_none opt_morphs ("Rep_" ^ name, "Abs_" ^ name); + val (Rep_name, Abs_name) = getOpt (opt_morphs, ("Rep_" ^ name, "Abs_" ^ name)); val setC = Const (full_name, setT); val RepC = Const (full Rep_name, newT --> oldT); val AbsC = Const (full Abs_name, oldT --> newT); @@ -251,7 +251,7 @@ fun sane_typedef prep_term def opt_name typ set opt_morphs tac = gen_typedef prep_term def - (if_none opt_name (#1 typ)) typ set opt_morphs all_tac [] [] (SOME tac); + (getOpt (opt_name, #1 typ)) typ set opt_morphs all_tac [] [] (SOME tac); fun add_typedef_x name typ set names thms tac = #1 o gen_typedef read_term true name typ set NONE (Tactic.rtac exI 1) names thms tac; @@ -287,8 +287,8 @@ in SOME (gr'', Codegen.mk_app brack (Pretty.str id) ps) end; fun get_name (Type (tname, _)) = tname | get_name _ = ""; - fun lookup f T = if_none (apsome f (Symtab.lookup - (TypedefData.get thy, get_name T))) "" + fun lookup f T = getOpt (Option.map f (Symtab.lookup + (TypedefData.get thy, get_name T)), "") in (case strip_comb t of (Const (s, Type ("fun", [T, U])), ts) => @@ -310,7 +310,7 @@ (case Symtab.lookup (TypedefData.get thy, s) of NONE => NONE | SOME (newT as Type (tname, Us), oldT, Abs_name, Rep_name) => - if is_some (Codegen.get_assoc_type thy tname) then NONE else + if isSome (Codegen.get_assoc_type thy tname) then NONE else let val sg = sign_of thy; val Abs_id = Codegen.mk_const_id sg Abs_name; @@ -382,7 +382,7 @@ Scan.option (P.$$$ "morphisms" |-- P.!!! (P.name -- P.name)); fun mk_typedef_proof ((((((def, opt_name), (vs, t)), mx), A), morphs)) = - typedef_proof ((def, if_none opt_name (Syntax.type_name t mx)), (t, vs, mx), A, morphs); + typedef_proof ((def, getOpt (opt_name, Syntax.type_name t mx)), (t, vs, mx), A, morphs); val typedefP = OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)" K.thy_goal diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/arith_data.ML --- a/src/HOL/arith_data.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/arith_data.ML Thu Mar 03 12:43:01 2005 +0100 @@ -25,7 +25,7 @@ (*normal form of sums: Suc (... (Suc (a + (b + ...))))*) fun mk_norm_sum ts = - let val (ones, sums) = partition (equal one) ts in + let val (ones, sums) = List.partition (equal one) ts in funpow (length ones) HOLogic.mk_Suc (mk_sum sums) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/cladata.ML --- a/src/HOL/cladata.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/cladata.ML Thu Mar 03 12:43:01 2005 +0100 @@ -33,7 +33,7 @@ (*prevent substitution on bool*) fun hyp_subst_tac' i thm = if i <= Thm.nprems_of thm andalso Term.exists_Const (fn ("op =", Type (_, [T, _])) => T <> Type ("bool", []) | _ => false) - (Library.nth_elem (i - 1, Thm.prems_of thm)) then hyp_subst_tac i thm else no_tac thm; + (List.nth (Thm.prems_of thm, i - 1)) then hyp_subst_tac i thm else no_tac thm; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/eqrule_HOL_data.ML --- a/src/HOL/eqrule_HOL_data.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/eqrule_HOL_data.ML Thu Mar 03 12:43:01 2005 +0100 @@ -45,14 +45,14 @@ (case Term.head_of p of Const(a,_) => (case Library.assoc(pairs,a) of - SOME(rls) => flat (map atoms ([th] RL rls)) + SOME(rls) => List.concat (map atoms ([th] RL rls)) | NONE => [th]) | _ => [th]) | _ => [th]) in atoms end; val prep_meta_eq = - (mapfilter + (List.mapPartial mk_eq o (mk_atomize tranformation_pairs) o Drule.gen_all diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/ex/svc_funcs.ML --- a/src/HOL/ex/svc_funcs.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/ex/svc_funcs.ML Thu Mar 03 12:43:01 2005 +0100 @@ -44,11 +44,11 @@ fun toString t = let fun ue (Buildin(s, l)) = - "(" ^ s ^ (foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") " + "(" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") " | ue (Interp(s, l)) = - "{" ^ s ^ (foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ "} " + "{" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ "} " | ue (UnInterp(s, l)) = - "(" ^ s ^ (foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") " + "(" ^ s ^ (Library.foldl (fn (a, b) => a ^ " " ^ (ue b)) ("", l)) ^ ") " | ue (FalseExpr) = "FALSE " | ue (TrueExpr) = "TRUE " | ue (Int i) = (signedInt i) ^ " " @@ -243,7 +243,7 @@ val body_e = mt pos body (*evaluate now to assign into !nat_vars*) in - foldr add_nat_var (!nat_vars, body_e) + Library.foldr add_nat_var (!nat_vars, body_e) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/hologic.ML --- a/src/HOL/hologic.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/hologic.ML Thu Mar 03 12:43:01 2005 +0100 @@ -148,7 +148,7 @@ fun all_const T = Const ("All", [T --> boolT] ---> boolT); fun mk_all (x, T, P) = all_const T $ absfree (x, T, P); -val list_all = foldr (fn ((x, T), P) => all_const T $ Abs (x, T, P)); +val list_all = Library.foldr (fn ((x, T), P) => all_const T $ Abs (x, T, P)); fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT); fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P); @@ -237,7 +237,7 @@ (*Makes a nested tuple from a list, following the product type structure*) fun mk_tuple (Type ("*", [T1, T2])) tms = mk_prod (mk_tuple T1 tms, - mk_tuple T2 (drop (length (prodT_factors T1), tms))) + mk_tuple T2 (Library.drop (length (prodT_factors T1), tms))) | mk_tuple T (t::_) = t; @@ -246,13 +246,13 @@ local (*currently unused*) -fun mk_tupleT Ts = foldr mk_prodT (Ts, unitT); +fun mk_tupleT Ts = Library.foldr mk_prodT (Ts, unitT); fun dest_tupleT (Type ("Product_Type.unit", [])) = [] | dest_tupleT (Type ("*", [T, U])) = T :: dest_tupleT U | dest_tupleT T = raise TYPE ("dest_tupleT", [T], []); -fun mk_tuple ts = foldr mk_prod (ts, unit); +fun mk_tuple ts = Library.foldr mk_prod (ts, unit); fun dest_tuple (Const ("Product_Type.Unity", _)) = [] | dest_tuple (Const ("Pair", _) $ t $ u) = t :: dest_tuple u diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/simpdata.ML --- a/src/HOL/simpdata.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/simpdata.ML Thu Mar 03 12:43:01 2005 +0100 @@ -253,14 +253,14 @@ (case head_of p of Const(a,_) => (case assoc(pairs,a) of - SOME(rls) => flat (map atoms ([th] RL rls)) + SOME(rls) => List.concat (map atoms ([th] RL rls)) | NONE => [th]) | _ => [th]) | _ => [th]) in atoms end; fun mksimps pairs = - (mapfilter (try mk_eq) o mk_atomize pairs o gen_all); + (List.mapPartial (try mk_eq) o mk_atomize pairs o gen_all); fun unsafe_solver_tac prems = FIRST'[resolve_tac(reflexive_thm::TrueI::refl::prems), atac, etac FalseE]; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOL/thy_syntax.ML --- a/src/HOL/thy_syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOL/thy_syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -14,7 +14,7 @@ fun mk_typedef_decl (((((opt_name, vs), t), mx), rhs), wt) = let - val name' = if_none opt_name t; + val name' = getOpt (opt_name,t); val name = unenclose name'; in (cat_lines [name', mk_triple (t, mk_list vs, mx), rhs, wt], @@ -101,11 +101,11 @@ fun mk_dt_struct (s, (id, i)) = s ^ "structure " ^ id ^ " =\n\ \struct\n\ - \ val distinct = nth_elem (" ^ i ^ ", distinct);\n\ - \ val inject = nth_elem (" ^ i ^ ", inject);\n\ - \ val exhaust = nth_elem (" ^ i ^ ", exhaustion);\n\ - \ val cases = nth_elem (" ^ i ^ ", case_thms);\n\ - \ val (split, split_asm) = nth_elem (" ^ i ^ ", split_thms);\n" ^ + \ val distinct = List.nth (distinct, " ^ i ^ ");\n\ + \ val inject = List.nth (inject, " ^ i ^ ");\n\ + \ val exhaust = List.nth (exhaustion, " ^ i ^ ");\n\ + \ val cases = List.nth (case_thms, " ^ i ^ ");\n\ + \ val (split, split_asm) = List.nth (split_thms, " ^ i ^ ");\n" ^ (if length names = 1 then " val induct = induction;\n\ \ val recs = rec_thms;\n\ @@ -114,7 +114,7 @@ else "") ^ "end;\n"; - val structs = foldl mk_dt_struct + val structs = Library.foldl mk_dt_struct ("", (names ~~ (map string_of_int (0 upto length names - 1)))); in @@ -131,7 +131,7 @@ fun mk_dt_string dts = let val names = map (fn ((((alt_name, _), name), _), _) => - unenclose (if_none alt_name name)) dts; + unenclose (getOpt (alt_name,name))) dts; val add_datatype_args = brackets (commas (map Library.quote names)) ^ " " ^ brackets (commas (map (fn ((((_, vs), name), mx), cs) => diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/IOA/Modelcheck/MuIOAOracle.ML --- a/src/HOLCF/IOA/Modelcheck/MuIOAOracle.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/IOA/Modelcheck/MuIOAOracle.ML Thu Mar 03 12:43:01 2005 +0100 @@ -2,7 +2,7 @@ (* call_sim_tac invokes oracle "Sim" *) fun call_sim_tac thm_list i state = let val sign = #sign (rep_thm state); - val (subgoal::_) = drop(i-1,prems_of state); + val (subgoal::_) = Library.drop(i-1,prems_of state); val OraAss = invoke_oracle MuIOAOracle.thy "Sim" (sign,SimOracleExn (subgoal,thm_list)); in (cut_facts_tac [OraAss] i) state @@ -13,7 +13,7 @@ fun is_sim_tac thm_list i state = let val sign = #sign (rep_thm state); in -case drop(i-1,prems_of state) of +case Library.drop(i-1,prems_of state) of [] => PureGeneral.Seq.empty | subgoal::_ => EVERY[REPEAT(etac thin_rl i), simp_tac (simpset() addsimps [ioa_implements_def]) i, diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/IOA/meta_theory/Sequence.ML --- a/src/HOLCF/IOA/meta_theory/Sequence.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/IOA/meta_theory/Sequence.ML Thu Mar 03 12:43:01 2005 +0100 @@ -1012,7 +1012,7 @@ fun all_frees_tac x i thm = let val tsig = Sign.tsig_of (Thm.sign_of_thm thm); - val frees = add_frees tsig (nth_elem (i - 1, prems_of thm), [x]); + val frees = add_frees tsig (List.nth (prems_of thm, i - 1), [x]); val frees' = sort (op >) (frees \ x) @ [x]; in foldl (qnt_tac i) (all_tac, frees') thm diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/adm.ML --- a/src/HOLCF/adm.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/adm.ML Thu Mar 03 12:43:01 2005 +0100 @@ -40,7 +40,7 @@ if i = lev then [[(Bound 0, path)]] else [] | find_subterms (t as (Abs (_, _, t2))) lev path = - if filter (fn x => x<=lev) + if List.filter (fn x => x<=lev) (add_loose_bnos (t, 0, [])) = [lev] then [(incr_bv (~lev, 0, t), path)]:: (find_subterms t2 (lev+1) (0::path)) @@ -52,7 +52,7 @@ | combine (x::xs) ys = (map (fn z => x @ z) ys) @ (combine xs ys) in - (if filter (fn x => x<=lev) + (if List.filter (fn x => x<=lev) (add_loose_bnos (t, 0, [])) = [lev] then [[(incr_bv (~lev, 0, t), path)]] else []) @ @@ -116,7 +116,7 @@ val tsig = Sign.tsig_of sign; val parTs = map snd (rev params); val rule = lift_rule (state, i) adm_subst; - val types = the o (fst (types_sorts rule)); + val types = valOf o (fst (types_sorts rule)); val tT = types ("t", j); val PT = types ("P", j); fun mk_abs [] t = t @@ -135,7 +135,7 @@ (*** extract subgoal i from proof state ***) -fun nth_subgoal i thm = nth_elem (i-1, prems_of thm); +fun nth_subgoal i thm = List.nth (prems_of thm, i-1); (*** the admissibility tactic ***) @@ -155,9 +155,9 @@ val prems = Logic.strip_assums_hyp goali; val params = Logic.strip_params goali; val ts = find_subterms t 0 []; - val ts' = filter eq_terms ts; - val ts'' = filter (is_chfin sign T params) ts'; - val thms = foldl (prove_cont tac sign s T prems params) ([], ts''); + val ts' = List.filter eq_terms ts; + val ts'' = List.filter (is_chfin sign T params) ts'; + val thms = Library.foldl (prove_cont tac sign s T prems params) ([], ts''); in (case thms of ((ts as ((t', _)::_), cont_thm)::_) => diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/cont_consts.ML --- a/src/HOLCF/cont_consts.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/cont_consts.ML Thu Mar 03 12:43:01 2005 +0100 @@ -43,7 +43,7 @@ val vnames = argnames (ord "A") n; val extra_parse_rule = Syntax.ParseRule (Constant name2, Constant name1); in [Syntax.ParsePrintRule (Syntax.mk_appl (Constant name2) (map Variable vnames), - foldl (fn (t,arg) => (Syntax.mk_appl (Constant "Rep_CFun") + Library.foldl (fn (t,arg) => (Syntax.mk_appl (Constant "Rep_CFun") [t,Variable arg])) (Constant name1,vnames))] @(case mx of InfixName _ => [extra_parse_rule] @@ -94,7 +94,7 @@ |> Theory.add_consts_i normal_decls |> Theory.add_consts_i (map first transformed_decls) |> Theory.add_syntax_i (map second transformed_decls) - |> Theory.add_trrules_i (flat (map third transformed_decls)) + |> Theory.add_trrules_i (List.concat (map third transformed_decls)) end; val add_consts = gen_add_consts Thm.read_ctyp; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/domain/axioms.ML --- a/src/HOLCF/domain/axioms.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/domain/axioms.ML Thu Mar 03 12:43:01 2005 +0100 @@ -31,7 +31,7 @@ val rep_iso_ax = ("rep_iso" ,mk_trp(dc_abs`(dc_rep`%x_name')=== %:x_name')); val when_def = ("when_def",%%:(dname^"_when") == - foldr (uncurry /\ ) (when_funs cons, /\x_name'((when_body cons (fn (x,y) => + Library.foldr (uncurry /\ ) (when_funs cons, /\x_name'((when_body cons (fn (x,y) => Bound(1+length cons+x-y)))`(dc_rep`Bound 0)))); fun con_def outer recu m n (_,args) = let @@ -43,10 +43,10 @@ fun inj y 1 _ = y | inj y _ 0 = %%:"sinl"`y | inj y i j = %%:"sinr"`(inj y (i-1) (j-1)); - in foldr /\# (args, outer (inj (parms args) m n)) end; + in Library.foldr /\# (args, outer (inj (parms args) m n)) end; val copy_def = ("copy_def", %%:(dname^"_copy") == /\"f" (dc_abs oo - foldl (op `) (%%:(dname^"_when") , + Library.foldl (op `) (%%:(dname^"_when") , mapn (con_def Id true (length cons)) 0 cons))); (* -- definitions concerning the constructors, discriminators and selectors - *) @@ -57,7 +57,7 @@ val dis_defs = let fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == mk_cRep_CFun(%%:(dname^"_when"),map - (fn (con',args) => (foldr /\# + (fn (con',args) => (Library.foldr /\# (args,if con'=con then %%:"TT" else %%:"FF"))) cons)) in map ddef cons end; @@ -65,8 +65,8 @@ fun sdef con n arg = (sel_of arg^"_def",%%:(sel_of arg) == mk_cRep_CFun(%%:(dname^"_when"),map (fn (con',args) => if con'<>con then %%:"UU" else - foldr /\# (args,Bound (length args - n))) cons)); - in flat(map (fn (con,args) => mapn (sdef con) 1 args) cons) end; + Library.foldr /\# (args,Bound (length args - n))) cons)); + in List.concat(map (fn (con,args) => mapn (sdef con) 1 args) cons) end; (* ----- axiom and definitions concerning induction ------------------------- *) @@ -100,7 +100,7 @@ let fun one_con (con,args) = let val nonrec_args = filter_out is_rec args; - val rec_args = filter is_rec args; + val rec_args = List.filter is_rec args; val recs_cnt = length rec_args; val allargs = nonrec_args @ rec_args @ map (upd_vname (fn s=> s^"'")) rec_args; @@ -114,10 +114,10 @@ (allargs~~((allargs_cnt-1) downto 0))); fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); - val capps = foldr mk_conj (mapn rel_app 1 rec_args, mk_conj( + val capps = Library.foldr mk_conj (mapn rel_app 1 rec_args, mk_conj( Bound(allargs_cnt+1)===mk_cRep_CFun(%%:con,map (bound_arg allvns) vns1), Bound(allargs_cnt+0)===mk_cRep_CFun(%%:con,map (bound_arg allvns) vns2))); - in foldr mk_ex (allvns, foldr mk_conj + in Library.foldr mk_ex (allvns, Library.foldr mk_conj (map (defined o Bound) nonlazy_idxs,capps)) end; fun one_comp n (_,cons) =mk_all(x_name(n+1),mk_all(x_name(n+1)^"'",mk_imp( proj (Bound 2) eqs n $ Bound 1 $ Bound 0, @@ -129,7 +129,7 @@ |> add_axioms_i (infer_types thy' dfs)(*add_defs_i*) |> add_axioms_i (infer_types thy' axs) |> Theory.parent_path; - val thy = foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs); + val thy = Library.foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs); in thy |> Theory.add_path comp_dnam |> add_axioms_i (infer_types thy' (bisim_def::(if length eqs>1 then [copy_def] else []))) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/domain/extender.ML --- a/src/HOLCF/domain/extender.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/domain/extender.ML Thu Mar 03 12:43:01 2005 +0100 @@ -44,11 +44,11 @@ val defaultS = Sign.defaultS sg; val test_dupl_typs = (case duplicates (map fst dtnvs) of [] => false | dups => error ("Duplicate types: " ^ commas_quote dups)); - val test_dupl_cons = (case duplicates (map first (flat cons'')) of + val test_dupl_cons = (case duplicates (map first (List.concat cons'')) of [] => false | dups => error ("Duplicate constructors: " ^ commas_quote dups)); val test_dupl_sels = (case duplicates - (map second (flat (map third (flat cons'')))) of + (map second (List.concat (map third (List.concat cons'')))) of [] => false | dups => error("Duplicate selectors: "^commas_quote dups)); val test_dupl_tvars = exists(fn s=>case duplicates(map(fst o rep_TFree)s)of [] => false | dups => error("Duplicate type arguments: " @@ -110,7 +110,7 @@ val eqs' = check_and_sort_domain (dtnvs,cons'') sg''; val thy' = thy'' |> Domain_Syntax.add_syntax (comp_dnam,eqs'); val dts = map (Type o fst) eqs'; - fun strip ss = drop (find_index_eq "'" ss +1, ss); + fun strip ss = Library.drop (find_index_eq "'" ss +1, ss); fun typid (Type (id,_)) = let val c = hd (Symbol.explode (Sign.base_name id)) in if Symbol.is_letter c then c else "t" end @@ -131,7 +131,7 @@ in theorems_thy |> Theory.add_path (Sign.base_name comp_dnam) - |> (#1 o (PureThy.add_thmss [(("rews", flat rewss @ take_rews), [])])) + |> (#1 o (PureThy.add_thmss [(("rews", List.concat rewss @ take_rews), [])])) |> Theory.parent_path end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/domain/interface.ML --- a/src/HOLCF/domain/interface.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/domain/interface.ML Thu Mar 03 12:43:01 2005 +0100 @@ -33,7 +33,7 @@ cat_lines (map (gen_vals dname) axioms1)^"\n"^ gen_vall "con_defs" (map (fn (con,_,_) => get dname (strip_esc con^"_def")) cons')^"\n"^ - gen_vall "sel_defs" (flat (map (fn (_,_,args) => map (fn (_,sel,_) => + gen_vall "sel_defs" (List.concat (map (fn (_,_,args) => map (fn (_,sel,_) => get dname (sel^"_def")) args) cons'))^"\n"^ gen_vall "dis_defs" (map (fn (con,_,_) => get dname (dis_name_ con^"_def")) cons')^"\n"^ diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/domain/library.ML --- a/src/HOLCF/domain/library.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/domain/library.ML Thu Mar 03 12:43:01 2005 +0100 @@ -13,12 +13,12 @@ fun mapn f n [] = [] | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; -fun foldr'' f (l,f2) = let fun itr [] = raise LIST "foldr''" +fun foldr'' f (l,f2) = let fun itr [] = raise Fail "foldr''" | itr [a] = f2 a | itr (a::l) = f(a, itr l) in itr l end; fun foldr' f l = foldr'' f (l,Id); -fun map_cumulr f start xs = foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => +fun map_cumulr f start xs = Library.foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => (y::ys,res2)) (xs,([],start)); @@ -98,7 +98,7 @@ fun is_rec arg = rec_of arg >=0; fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); fun nonlazy args = map vname (filter_out is_lazy args); -fun nonlazy_rec args = map vname (filter is_nonlazy_rec args); +fun nonlazy_rec args = map vname (List.filter is_nonlazy_rec args); (* ----- support for type and mixfix expressions ----- *) @@ -123,7 +123,7 @@ | sg _ = Imposs "library:sg"; fun sf [] = %%:"_emptysort" | sf s = %%:"_sort" $ sg s - fun tf (Type (s,args)) = foldl (op $) (%:s,map tf args) + fun tf (Type (s,args)) = Library.foldl (op $) (%:s,map tf args) | tf (TFree(s,sort)) = %%:"_ofsort" $ %:s $ sf sort | tf _ = Imposs "library:tf"; in @@ -146,7 +146,7 @@ infix 9 ` ; fun f` x = %%:"Rep_CFun" $ f $ x; infix 9 `% ; fun f`% s = f` %: s; infix 9 `%%; fun f`%%s = f` %%:s; -fun mk_cRep_CFun (F,As) = foldl (op `) (F,As); +fun mk_cRep_CFun (F,As) = Library.foldl (op `) (F,As); fun con_app2 con f args = mk_cRep_CFun(%%:con,map f args); fun con_app con = con_app2 con %#; fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; @@ -162,7 +162,7 @@ (fn S => fn t => Const("cfst",mk_prodT(dummyT,t)->>dummyT)`S) (fn S => fn t => Const("csnd",mk_prodT(t,dummyT)->>dummyT)`S) T (map ((fn tp => tp ->> tp) o Type o fix_tp o fst) eqs); -fun lift tfn = foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); +fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); fun /\ v T = %%:"Abs_CFun" $ mk_lam(v,T); fun /\# (arg,T) = /\ (vname arg) T; diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/domain/syntax.ML --- a/src/HOLCF/domain/syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/domain/syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -22,14 +22,14 @@ else foldr' mk_sprodT (map opt_lazy args); fun freetvar s = let val tvar = mk_TFree s in if tvar mem typevars then freetvar ("t"^s) else tvar end; - fun when_type (_ ,_,args) = foldr (op ->>) (map third args,freetvar "t"); + fun when_type (_ ,_,args) = Library.foldr (op ->>) (map third args,freetvar "t"); in val dtype = Type(dname,typevars); val dtype2 = foldr' mk_ssumT (map prod cons'); val dnam = Sign.base_name dname; val const_rep = (dnam^"_rep" , dtype ->> dtype2, NoSyn); val const_abs = (dnam^"_abs" , dtype2 ->> dtype , NoSyn); - val const_when = (dnam^"_when",foldr (op ->>) ((map when_type cons'), + val const_when = (dnam^"_when",Library.foldr (op ->>) ((map when_type cons'), dtype ->> freetvar "t"), NoSyn); val const_copy = (dnam^"_copy", dtypeprod ->> dtype ->> dtype , NoSyn); end; @@ -42,7 +42,7 @@ else c::esc cs | esc [] = [] in implode o esc o Symbol.explode end; - fun con (name,s,args) = (name,foldr (op ->>) (map third args,dtype),s); + fun con (name,s,args) = (name,Library.foldr (op ->>) (map third args,dtype),s); fun dis (con ,s,_ ) = (dis_name_ con, dtype->>trT, Mixfix(escape ("is_" ^ con), [], Syntax.max_pri)); (* stricly speaking, these constants have one argument, @@ -52,7 +52,7 @@ in val consts_con = map con cons'; val consts_dis = map dis cons'; - val consts_sel = flat(map sel cons'); + val consts_sel = List.concat(map sel cons'); end; (* ----- constants concerning induction ------------------------------------- *) @@ -70,7 +70,7 @@ (string_of_int m)); fun app s (l,r) = mk_appl (Constant s) [l,r]; fun case1 n (con,mx,args) = mk_appl (Constant "_case1") - [foldl (app "Rep_CFun") (c_ast con mx, (mapn (argvar n) 1 args)), + [Library.foldl (app "Rep_CFun") (c_ast con mx, (mapn (argvar n) 1 args)), expvar n]; fun arg1 n (con,_,args) = if args = [] then expvar n else mk_appl (Constant "LAM ") @@ -80,7 +80,7 @@ (mk_appl (Constant "_case_syntax") [Variable "x", foldr' (fn (c,cs) => mk_appl (Constant"_case2") [c,cs]) (mapn case1 1 cons')], - mk_appl (Constant "Rep_CFun") [foldl + mk_appl (Constant "Rep_CFun") [Library.foldl (fn (w,a ) => mk_appl (Constant"Rep_CFun" ) [w,a ]) (Constant (dnam^"_when"),mapn arg1 1 cons'), Variable "x"]) @@ -107,10 +107,10 @@ val const_copy = (comp_dnam^"_copy" ,funprod ->> funprod, NoSyn); val const_bisim = (comp_dnam^"_bisim" ,relprod --> boolT , NoSyn); val ctt = map (calc_syntax funprod) eqs'; -in thy'' |> ContConsts.add_consts_i (flat (map fst ctt) @ +in thy'' |> ContConsts.add_consts_i (List.concat (map fst ctt) @ (if length eqs'>1 then [const_copy] else[])@ [const_bisim]) - |> Theory.add_trrules_i (flat(map snd ctt)) + |> Theory.add_trrules_i (List.concat(map snd ctt)) end; (* let *) end; (* local *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/HOLCF/domain/theorems.ML --- a/src/HOLCF/domain/theorems.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/HOLCF/domain/theorems.ML Thu Mar 03 12:43:01 2005 +0100 @@ -71,7 +71,7 @@ val ax_when_def = ga "when_def" dname; val axs_con_def = map (fn (con,_) => ga (extern_name con^"_def") dname) cons; val axs_dis_def = map (fn (con,_) => ga ( dis_name con^"_def") dname) cons; -val axs_sel_def = flat(map (fn (_,args) => +val axs_sel_def = List.concat(map (fn (_,args) => map (fn arg => ga (sel_of arg ^"_def") dname) args) cons); val ax_copy_def = ga "copy_def" dname; @@ -101,7 +101,7 @@ etac (ax_rep_iso RS subst) 1]; fun exh foldr1 cn quant foldr2 var = let fun one_con (con,args) = let val vns = map vname args in - foldr quant (vns, foldr2 ((%:x_name === con_app2 con (var vns) vns):: + Library.foldr quant (vns, foldr2 ((%:x_name === con_app2 con (var vns) vns):: map (defined o (var vns)) (nonlazy args))) end in foldr1 ((cn(%:x_name===UU))::map one_con cons) end; in @@ -130,9 +130,9 @@ prod_tac args THEN sum_rest_tac p) THEN sum_tac cons' prems | sum_tac _ _ = Imposs "theorems:sum_tac"; - in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%:"P"))) + in pg'' thy [] (exh (fn l => Library.foldr (op ===>) (l,mk_trp(%:"P"))) (fn T => T ==> %:"P") mk_All - (fn l => foldr (op ===>) (map mk_trp l, + (fn l => Library.foldr (op ===>) (map mk_trp l, mk_trp(%:"P"))) bound_arg) (fn prems => [ @@ -153,9 +153,9 @@ end; local - fun bind_fun t = foldr mk_All (when_funs cons,t); + fun bind_fun t = Library.foldr mk_All (when_funs cons,t); fun bound_fun i _ = Bound (length cons - i); - val when_app = foldl (op `) (%%:(dname^"_when"), mapn bound_fun 1 cons); + val when_app = Library.foldl (op `) (%%:(dname^"_when"), mapn bound_fun 1 cons); val when_appl = pg [ax_when_def] (bind_fun(mk_trp(when_app`%x_name === when_body cons (fn (m,n)=> bound_fun (n-m) 0)`(dc_rep`%x_name))))[ simp_tac HOLCF_ss 1]; @@ -182,7 +182,7 @@ (mk_trp((%%:(dis_name c))`(con_app con args) === %%:(if con=c then "TT" else "FF"))))) [ asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in flat(map (fn (c,_) => map (one_dis c) cons) cons) end; + in List.concat(map (fn (c,_) => map (one_dis c) cons) cons) end; val dis_defins = map (fn (con,args) => pg [] (defined(%:x_name) ==> defined(%%:(dis_name con)`%x_name)) [ rtac casedist 1, @@ -191,7 +191,7 @@ (HOLCF_ss addsimps dis_apps) 1))]) cons; in dis_stricts @ dis_defins @ dis_apps end; -val con_stricts = flat(map (fn (con,args) => map (fn vn => +val con_stricts = List.concat(map (fn (con,args) => map (fn vn => pg (axs_con_def) (mk_trp(con_app2 con (fn arg => if vname arg = vn then UU else %# arg) args === UU))[ @@ -207,22 +207,22 @@ val sel_stricts = let fun one_sel sel = pg axs_sel_def (mk_trp(strict(%%:sel))) [ simp_tac (HOLCF_ss addsimps when_rews) 1]; -in flat(map (fn (_,args) =>map (fn arg => one_sel (sel_of arg)) args) cons) end; +in List.concat(map (fn (_,args) =>map (fn arg => one_sel (sel_of arg)) args) cons) end; val sel_apps = let fun one_sel c n sel = map (fn (con,args) => let val nlas = nonlazy args; val vns = map vname args; in pg axs_sel_def (lift_defined %: - (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas, + (List.filter (fn v => con=c andalso (v<>List.nth(vns,n))) nlas, mk_trp((%%:sel)`(con_app con args) === - (if con=c then %:(nth_elem(n,vns)) else UU)))) + (if con=c then %:(List.nth(vns,n)) else UU)))) ( (if con=c then [] else map(case_UU_tac(when_rews@con_stricts)1) nlas) - @(if con=c andalso ((nth_elem(n,vns)) mem nlas) + @(if con=c andalso ((List.nth(vns,n)) mem nlas) then[case_UU_tac (when_rews @ con_stricts) 1 - (nth_elem(n,vns))] else []) + (List.nth(vns,n))] else []) @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons; -in flat(map (fn (c,args) => - flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end; +in List.concat(map (fn (c,args) => + List.concat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end; val sel_defins = if length cons=1 then map (fn arg => pg [](defined(%:x_name)==> defined(%%:(sel_of arg)`%x_name)) [ rtac casedist 1, @@ -249,7 +249,7 @@ fun distincts [] = [] | distincts (c::cs) = (map (distinct c) cs) :: distincts cs; in distincts cons end; -val dist_les = flat (flat distincts_le); +val dist_les = List.concat (List.concat distincts_le); val dist_eqs = let fun distinct (_,args1) ((_,args2),leqs) = let val (le1,le2) = (hd leqs, hd(tl leqs)); @@ -273,10 +273,10 @@ mk_trp (foldr' mk_conj (ListPair.map rel (map %# largs, map %# rargs)))))) end; - val cons' = filter (fn (_,args) => args<>[]) cons; + val cons' = List.filter (fn (_,args) => args<>[]) cons; in val inverts = map (fn (con,args) => - pgterm (op <<) con args (flat(map (fn arg => [ + pgterm (op <<) con args (List.concat(map (fn arg => [ TRY(rtac conjI 1), dres_inst_tac [("fo",sel_of arg)] monofun_cfun_arg 1, asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1] @@ -302,7 +302,7 @@ (con_app2 con (app_rec_arg (cproj (%:"f") eqs)) args)))) (map (case_UU_tac (abs_strict::when_strict::con_stricts) 1 o vname) - (filter (fn a => not (is_rec a orelse is_lazy a)) args) + (List.filter (fn a => not (is_rec a orelse is_lazy a)) args) @[asm_simp_tac (HOLCF_ss addsimps when_apps) 1, simp_tac (HOLCF_ss addsimps axs_con_def) 1]))cons; val copy_stricts = map (fn (con,args) => pg [] (mk_trp(dc_copy`UU` @@ -310,7 +310,7 @@ (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews in map (case_UU_tac rews 1) (nonlazy args) @ [ asm_simp_tac (HOLCF_ss addsimps rews) 1] end)) - (filter (fn (_,args)=>exists is_nonlazy_rec args) cons); + (List.filter (fn (_,args)=>exists is_nonlazy_rec args) cons); val copy_rews = copy_strict::copy_apps @ copy_stricts; in thy |> Theory.add_path (Sign.base_name dname) |> (#1 o (PureThy.add_thmss (map Thm.no_attributes [ @@ -352,8 +352,8 @@ local fun gt s dn = get_thm thy (dn ^ "." ^ s, NONE); fun gts s dn = get_thms thy (dn ^ "." ^ s, NONE) in val cases = map (gt "casedist" ) dnames; -val con_rews = flat (map (gts "con_rews" ) dnames); -val copy_rews = flat (map (gts "copy_rews") dnames); +val con_rews = List.concat (map (gts "con_rews" ) dnames); +val copy_rews = List.concat (map (gts "copy_rews") dnames); end; (* local *) fun dc_take dn = %%:(dn^"_take"); @@ -379,17 +379,17 @@ simp_tac iterate_Cprod_ss 1]) 1 dnames; val c_UU_tac = case_UU_tac (take_stricts'::copy_con_rews) 1; val take_apps = pg copy_take_defs (mk_trp(foldr' mk_conj - (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all + (List.concat(map (fn ((dn,_),cons) => map (fn (con,args) => Library.foldr mk_all (map vname args,(dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args) === - con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %:"n")) + con_app2 con (app_rec_arg (fn n=>dc_take (List.nth(dnames,n))$ %:"n")) args)) cons) eqs)))) ([ simp_tac iterate_Cprod_ss 1, induct_tac "n" 1, simp_tac(iterate_Cprod_ss addsimps copy_con_rews) 1, asm_full_simp_tac (HOLCF_ss addsimps - (filter (has_fewer_prems 1) copy_rews)) 1, + (List.filter (has_fewer_prems 1) copy_rews)) 1, TRY(safe_tac HOL_cs)] @ - (flat(map (fn ((dn,_),cons) => map (fn (con,args) => + (List.concat(map (fn ((dn,_),cons) => map (fn (con,args) => if nonlazy_rec args = [] then all_tac else EVERY(map c_UU_tac (nonlazy_rec args)) THEN asm_full_simp_tac (HOLCF_ss addsimps copy_rews)1 @@ -399,23 +399,23 @@ end; (* local *) local - fun one_con p (con,args) = foldr mk_All (map vname args, + fun one_con p (con,args) = Library.foldr mk_All (map vname args, lift_defined (bound_arg (map vname args)) (nonlazy args, lift (fn arg => %:(P_name (1+rec_of arg)) $ bound_arg args arg) - (filter is_rec args,mk_trp(%:p $ con_app2 con (bound_arg args) args)))); + (List.filter is_rec args,mk_trp(%:p $ con_app2 con (bound_arg args) args)))); fun one_eq ((p,cons),concl) = (mk_trp(%:p $ UU) ===> - foldr (op ===>) (map (one_con p) cons,concl)); - fun ind_term concf = foldr one_eq (mapn (fn n => fn x => (P_name n, x))1conss, + Library.foldr (op ===>) (map (one_con p) cons,concl)); + fun ind_term concf = Library.foldr one_eq (mapn (fn n => fn x => (P_name n, x))1conss, mk_trp(foldr' mk_conj (mapn concf 1 dnames))); val take_ss = HOL_ss addsimps take_rews; fun quant_tac i = EVERY(mapn(fn n=> fn _=> res_inst_tac[("x",x_name n)]spec i) 1 dnames); - fun ind_prems_tac prems = EVERY(flat (map (fn cons => ( + fun ind_prems_tac prems = EVERY(List.concat (map (fn cons => ( resolve_tac prems 1 :: - flat (map (fn (_,args) => + List.concat (map (fn (_,args) => resolve_tac prems 1 :: map (K(atac 1)) (nonlazy args) @ - map (K(atac 1)) (filter is_rec args)) + map (K(atac 1)) (List.filter is_rec args)) cons))) conss)); local (* check whether every/exists constructor of the n-th part of the equation: @@ -425,11 +425,11 @@ is_rec arg andalso not(rec_of arg mem ns) andalso ((rec_of arg = n andalso nfn(lazy_rec orelse is_lazy arg)) orelse rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) - (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) + (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg)))) ) o snd) cons; fun all_rec_to ns = rec_to forall not all_rec_to ns; fun warn (n,cons) = if all_rec_to [] false (n,cons) then (warning - ("domain "^nth_elem(n,dnames)^" is empty!"); true) else false; + ("domain "^List.nth(dnames,n)^" is empty!"); true) else false; fun lazy_rec_to ns = rec_to exists Id lazy_rec_to ns; in val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs; @@ -444,18 +444,18 @@ induct_tac "n" 1, simp_tac (take_ss addsimps prems) 1, TRY(safe_tac HOL_cs)] - @ flat(map (fn (cons,cases) => [ + @ List.concat(map (fn (cons,cases) => [ res_inst_tac [("x","x")] cases 1, asm_simp_tac (take_ss addsimps prems) 1] - @ flat(map (fn (con,args) => + @ List.concat(map (fn (con,args) => asm_simp_tac take_ss 1 :: map (fn arg => case_UU_tac (prems@con_rews) 1 ( - nth_elem(rec_of arg,dnames)^"_take n$"^vname arg)) - (filter is_nonlazy_rec args) @ [ + List.nth(dnames,rec_of arg)^"_take n$"^vname arg)) + (List.filter is_nonlazy_rec args) @ [ resolve_tac prems 1] @ map (K (atac 1)) (nonlazy args) @ - map (K (etac spec 1)) (filter is_rec args)) + map (K (etac spec 1)) (List.filter is_rec args)) cons)) (conss~~cases))); @@ -496,14 +496,14 @@ induct_tac "n" 1, simp_tac take_ss 1, TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @ - flat(mapn (fn n => fn (cons,cases) => [ + List.concat(mapn (fn n => fn (cons,cases) => [ simp_tac take_ss 1, rtac allI 1, res_inst_tac [("x",x_name n)] cases 1, asm_simp_tac take_ss 1] @ - flat(map (fn (con,args) => + List.concat(map (fn (con,args) => asm_simp_tac take_ss 1 :: - flat(map (fn vn => [ + List.concat(map (fn vn => [ eres_inst_tac [("x",vn)] all_dupE 1, etac disjE 1, asm_simp_tac (HOL_ss addsimps con_rews) 1, @@ -523,7 +523,7 @@ (finites, pg'' thy[](ind_term (fn n => fn dn => %:(P_name n) $ %:(x_name n)))(fn prems => TRY(safe_tac HOL_cs) :: - flat (map (fn (finite,fin_ind) => [ + List.concat (map (fn (finite,fin_ind) => [ rtac(rewrite_rule axs_finite_def finite RS exE)1, etac subst 1, rtac fin_ind 1, @@ -532,7 +532,7 @@ ) end (* let *) else (mapn (fn n => fn dn => read_instantiate_sg (sign_of thy) [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames, - pg'' thy [] (foldr (op ===>) (mapn (fn n => K(mk_trp(%%:"adm" $ %:(P_name n)))) + pg'' thy [] (Library.foldr (op ===>) (mapn (fn n => K(mk_trp(%%:"adm" $ %:(P_name n)))) 1 dnames, ind_term (fn n => fn dn => %:(P_name n) $ %:(x_name n)))) (fn prems => map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [ @@ -556,8 +556,8 @@ val take_ss = HOL_ss addsimps take_rews; val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")")); val coind_lemma=pg[ax_bisim_def](mk_trp(mk_imp(%%:(comp_dname^"_bisim") $ %:"R", - foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs, - foldr mk_imp (mapn (fn n => K(proj (%:"R") eqs n $ + Library.foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs, + Library.foldr mk_imp (mapn (fn n => K(proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1)) 0 dnames, foldr' mk_conj (mapn (fn n => fn dn => (dc_take dn $ %:"n" `bnd_arg n 0 === @@ -566,7 +566,7 @@ induct_tac "n" 1, simp_tac take_ss 1, safe_tac HOL_cs] @ - flat(mapn (fn n => fn x => [ + List.concat(mapn (fn n => fn x => [ rotate_tac (n+1) 1, etac all2E 1, eres_inst_tac [("P1", sproj "R" eqs n^ @@ -576,11 +576,11 @@ 0 xs)); in val coind = pg [] (mk_trp(%%:(comp_dname^"_bisim") $ %:"R") ===> - foldr (op ===>) (mapn (fn n => fn x => + Library.foldr (op ===>) (mapn (fn n => fn x => mk_trp(proj (%:"R") eqs n $ %:x $ %:(x^"'"))) 0 xs, mk_trp(foldr' mk_conj (map (fn x => %:x === %:(x^"'")) xs)))) ([ TRY(safe_tac HOL_cs)] @ - flat(map (fn take_lemma => [ + List.concat(map (fn take_lemma => [ rtac take_lemma 1, cut_facts_tac [coind_lemma] 1, fast_tac HOL_cs 1]) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/Arith/abel_cancel.ML --- a/src/Provers/Arith/abel_cancel.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/Arith/abel_cancel.ML Thu Mar 03 12:43:01 2005 +0100 @@ -160,7 +160,7 @@ val _ = if null common then raise Cancel (*nothing to do*) else () - fun cancelled tms = mk_sum ty (foldl cancel1 (tms, common)) + fun cancelled tms = mk_sum ty (Library.foldl cancel1 (tms, common)) val lt' = cancelled ltms and rt' = cancelled rtms diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/Arith/cancel_numeral_factor.ML --- a/src/Provers/Arith/cancel_numeral_factor.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/Arith/cancel_numeral_factor.ML Thu Mar 03 12:43:01 2005 +0100 @@ -75,7 +75,7 @@ Data.prove_conv [Data.norm_tac] sg hyps xs (t', Data.mk_bal (newshape(n1,t1'), newshape(n2,t2'))) in - apsome Data.simplify_meta_eq + Option.map Data.simplify_meta_eq (Data.prove_conv [Data.trans_tac reshape, rtac Data.cancel 1, Data.numeral_simp_tac] sg hyps xs (t', rhs)) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/Arith/cancel_numerals.ML --- a/src/Provers/Arith/cancel_numerals.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/Arith/cancel_numerals.ML Thu Mar 03 12:43:01 2005 +0100 @@ -59,11 +59,11 @@ (*a left-to-right scan of terms1, seeking a term of the form #n*u, where #m*u is in terms2 for some m*) fun find_common (terms1,terms2) = - let val tab2 = foldl update_by_coeff (Termtab.empty, terms2) + let val tab2 = Library.foldl update_by_coeff (Termtab.empty, terms2) fun seek [] = raise TERM("find_common", []) | seek (t::terms) = let val (_,u) = Data.dest_coeff t - in if is_some (Termtab.lookup (tab2, u)) then u + in if isSome (Termtab.lookup (tab2, u)) then u else seek terms end in seek terms1 end; @@ -91,7 +91,7 @@ Data.mk_bal (newshape(n1,terms1'), newshape(n2,terms2'))) in - apsome Data.simplify_meta_eq + Option.map Data.simplify_meta_eq (if n2<=n1 then Data.prove_conv [Data.trans_tac reshape, rtac Data.bal_add1 1, diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/Arith/combine_numerals.ML --- a/src/Provers/Arith/combine_numerals.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/Arith/combine_numerals.ML Thu Mar 03 12:43:01 2005 +0100 @@ -78,7 +78,7 @@ Data.mk_sum T ([Data.mk_coeff(m,u), Data.mk_coeff(n,u)] @ terms)) in - apsome Data.simplify_meta_eq + Option.map Data.simplify_meta_eq (Data.prove_conv [Data.trans_tac reshape, rtac Data.left_distrib 1, Data.numeral_simp_tac] sg xs diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/Arith/extract_common_term.ML --- a/src/Provers/Arith/extract_common_term.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/Arith/extract_common_term.ML Thu Mar 03 12:43:01 2005 +0100 @@ -42,10 +42,10 @@ (*a left-to-right scan of terms1, seeking a term u that is also in terms2*) fun find_common (terms1,terms2) = - let val tab2 = foldl update_by_coeff (Termtab.empty, terms2) + let val tab2 = Library.foldl update_by_coeff (Termtab.empty, terms2) fun seek [] = raise TERM("find_common", []) | seek (u::terms) = - if is_some (Termtab.lookup (tab2, u)) then u + if isSome (Termtab.lookup (tab2, u)) then u else seek terms in seek terms1 end; @@ -68,7 +68,7 @@ Data.mk_bal (Data.mk_sum T (u::terms1'), Data.mk_sum T (u::terms2'))) in - apsome Data.simplify_meta_eq reshape + Option.map Data.simplify_meta_eq reshape end handle TERM _ => NONE | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/Arith/fast_lin_arith.ML --- a/src/Provers/Arith/fast_lin_arith.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/Arith/fast_lin_arith.ML Thu Mar 03 12:43:01 2005 +0100 @@ -171,7 +171,7 @@ (* PRE: ex[v] must be 0! *) fun eval (ex:rat list) v (a:int,le,cs:int list) = let val rs = map rat_of_int cs - val rsum = foldl ratadd (rat0,map ratmul (rs ~~ ex)) + val rsum = Library.foldl ratadd (rat0,map ratmul (rs ~~ ex)) in (ratmul(ratadd(rat_of_int a,ratneg rsum), ratinv(el v rs)), le) end; (* If el v rs < 0, le should be negated. Instead this swap is taken into account in ratrelmin2. @@ -217,14 +217,14 @@ in if ratle(lb,ub') then ub' else raise NoEx end; fun findex1 discr (ex,(v,lineqs)) = - let val nz = filter (fn (Lineq(_,_,cs,_)) => el v cs <> 0) lineqs; - val ineqs = foldl elim_eqns ([],nz) - val (ge,le) = partition (fn (_,_,cs) => el v cs > 0) ineqs + let val nz = List.filter (fn (Lineq(_,_,cs,_)) => el v cs <> 0) lineqs; + val ineqs = Library.foldl elim_eqns ([],nz) + val (ge,le) = List.partition (fn (_,_,cs) => el v cs > 0) ineqs val lb = ratrelmax(map (eval ex v) ge) val ub = ratrelmin(map (eval ex v) le) - in nth_update (choose2 (nth_elem(v,discr)) (lb,ub)) (v,ex) end; + in nth_update (choose2 (List.nth(discr,v)) (lb,ub)) (v,ex) end; -fun findex discr = foldl (findex1 discr); +fun findex discr = Library.foldl (findex1 discr); fun elim1 v x = map (fn (a,le,bs) => (ratadd(a,ratneg(ratmul(el v bs,x))), le, @@ -239,9 +239,9 @@ in case nz of [] => ex | (_,_,cs) :: _ => let val v = find_index (not o equal rat0) cs - val d = nth_elem(v,discr) + val d = List.nth(discr,v) val pos = ratge0(el v cs) - val sv = filter (single_var v) nz + val sv = List.filter (single_var v) nz val minmax = if pos then if d then ratroundup o fst o ratrelmax else ratexact true o ratrelmax @@ -255,7 +255,7 @@ end; fun findex0 discr n lineqs = - let val ineqs = foldl elim_eqns ([],lineqs) + let val ineqs = Library.foldl elim_eqns ([],lineqs) val rineqs = map (fn (a,le,cs) => (rat_of_int a, le, map rat_of_int cs)) ineqs in pick_vars discr (rineqs,replicate n rat0) end; @@ -321,7 +321,7 @@ case ty of Eq => k <> 0 | Le => k > 0 | Lt => k >= 0; fun calc_blowup l = - let val (p,n) = partition (apl(0,op<)) (filter (apl(0,op<>)) l) + let val (p,n) = List.partition (apl(0,op<)) (List.filter (apl(0,op<>)) l) in (length p) * (length n) end; (* ------------------------------------------------------------------------- *) @@ -337,7 +337,7 @@ (* ------------------------------------------------------------------------- *) fun allpairs f xs ys = - flat(map (fn x => map (fn y => f x y) ys) xs); + List.concat(map (fn x => map (fn y => f x y) ys) xs); fun extract_first p = let fun extract xs (y::ys) = if p y then (SOME y,xs@ys) @@ -358,7 +358,7 @@ fun elim(ineqs,hist) = let val dummy = print_ineqs ineqs; - val (triv,nontriv) = partition is_trivial ineqs in + val (triv,nontriv) = List.partition is_trivial ineqs in if not(null triv) then case Library.find_first is_answer triv of NONE => elim(nontriv,hist) @@ -366,16 +366,16 @@ else if null nontriv then Failure(hist) else - let val (eqs,noneqs) = partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in + let val (eqs,noneqs) = List.partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in if not(null eqs) then - let val clist = foldl (fn (cs,Lineq(_,_,l,_)) => l union cs) ([],eqs) + let val clist = Library.foldl (fn (cs,Lineq(_,_,l,_)) => l union cs) ([],eqs) val sclist = sort (fn (x,y) => int_ord(abs(x),abs(y))) - (filter (fn i => i<>0) clist) + (List.filter (fn i => i<>0) clist) val c = hd sclist val (SOME(eq as Lineq(_,_,ceq,_)),othereqs) = extract_first (fn Lineq(_,_,l,_) => c mem l) eqs val v = find_index_eq c ceq - val (ioth,roth) = partition (fn (Lineq(_,_,l,_)) => el v l = 0) + val (ioth,roth) = List.partition (fn (Lineq(_,_,l,_)) => el v l = 0) (othereqs @ noneqs) val others = map (elim_var v eq) roth @ ioth in elim(others,(v,nontriv)::hist) end @@ -385,12 +385,12 @@ val coeffs = map (fn i => map (el i) lists) numlist val blows = map calc_blowup coeffs val iblows = blows ~~ numlist - val nziblows = filter (fn (i,_) => i<>0) iblows + val nziblows = List.filter (fn (i,_) => i<>0) iblows in if null nziblows then Failure((~1,nontriv)::hist) else let val (c,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows) - val (no,yes) = partition (fn (Lineq(_,_,l,_)) => el v l = 0) ineqs - val (pos,neg) = partition(fn (Lineq(_,_,l,_)) => el v l > 0) yes + val (no,yes) = List.partition (fn (Lineq(_,_,l,_)) => el v l = 0) ineqs + val (pos,neg) = List.partition(fn (Lineq(_,_,l,_)) => el v l > 0) yes in elim(no @ allpairs (elim_var v) pos neg, (v,nontriv)::hist) end end end @@ -420,9 +420,9 @@ in fun mkthm sg asms just = let val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} = Data.get_sg sg; - val atoms = foldl (fn (ats,(lhs,_,_,rhs,_,_)) => + val atoms = Library.foldl (fn (ats,(lhs,_,_,rhs,_,_)) => map fst lhs union (map fst rhs union ats)) - ([], mapfilter (fn thm => if Thm.no_prems thm + ([], List.mapPartial (fn thm => if Thm.no_prems thm then LA_Data.decomp sg (concl_of thm) else NONE) asms) @@ -438,7 +438,7 @@ fun addthms thm1 thm2 = case add2 thm1 thm2 of NONE => (case try_add ([thm1] RL inj_thms) thm2 of - NONE => ( the(try_add ([thm2] RL inj_thms) thm1) + NONE => ( valOf(try_add ([thm2] RL inj_thms) thm1) handle OPTION => (trace_thm "" thm1; trace_thm "" thm2; sys_error "Lin.arith. failed to add thms") @@ -468,8 +468,8 @@ let val thm' = trace_thm "Simplified:" (full_simplify simpset thm) in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end - fun mk(Asm i) = trace_thm "Asm" (nth_elem(i,asms)) - | mk(Nat i) = (trace_msg "Nat"; LA_Logic.mk_nat_thm sg (nth_elem(i,atoms))) + fun mk(Asm i) = trace_thm "Asm" (List.nth(asms,i)) + | mk(Nat i) = (trace_msg "Nat"; LA_Logic.mk_nat_thm sg (List.nth(atoms,i))) | mk(LessD(j)) = trace_thm "L" (hd([mk j] RL lessD)) | mk(NotLeD(j)) = trace_thm "NLe" (mk j RS LA_Logic.not_leD) | mk(NotLeDD(j)) = trace_thm "NLeD" (hd([mk j RS LA_Logic.not_leD] RL lessD)) @@ -484,7 +484,7 @@ in trace_thm "After simplification:" fls; if LA_Logic.is_False fls then fls else - (tracing "Assumptions:"; seq print_thm asms; + (tracing "Assumptions:"; List.app print_thm asms; tracing "Proved:"; print_thm fls; warning "Linear arithmetic should have refuted the assumptions.\n\ \Please inform Tobias Nipkow (nipkow@in.tum.de)."; @@ -496,7 +496,7 @@ fun coeff poly atom = case assoc(poly,atom) of NONE => 0 | SOME i => i; -fun lcms is = foldl lcm (1,is); +fun lcms is = Library.foldl lcm (1,is); fun integ(rlhs,r,rel,rrhs,s,d) = let val (rn,rd) = rep_rat r and (sn,sd) = rep_rat s @@ -593,17 +593,17 @@ fun add_datoms(dats,((lhs,_,_,rhs,_,d),_)) = (map (pair d o fst) lhs) union ((map (pair d o fst) rhs) union dats) -fun discr initems = map fst (foldl add_datoms ([],initems)); +fun discr initems = map fst (Library.foldl add_datoms ([],initems)); fun refutes sg (pTs,params) ex = let fun refute (initems::initemss) js = - let val atoms = foldl add_atoms ([],initems) + let val atoms = Library.foldl add_atoms ([],initems) val n = length atoms val mkleq = mklineq n atoms val ixs = 0 upto (n-1) val iatoms = atoms ~~ ixs - val natlineqs = mapfilter (mknat pTs ixs) iatoms + val natlineqs = List.mapPartial (mknat pTs ixs) iatoms val ineqs = map mkleq initems @ natlineqs in case elim(ineqs,[]) of Success(j) => @@ -628,7 +628,7 @@ end state; -fun count P xs = length(filter P xs); +fun count P xs = length(List.filter P xs); (* The limit on the number of ~= allowed. Because each ~= is split into two cases, this can lead to an explosion. diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/blast.ML --- a/src/Provers/blast.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/blast.ML Thu Mar 03 12:43:01 2005 +0100 @@ -134,7 +134,7 @@ fun rand (f$x) = x; (* maps (f, [t1,...,tn]) to f(t1,...,tn) *) -val list_comb : term * term list -> term = foldl (op $); +val list_comb : term * term list -> term = Library.foldl (op $); (* maps f(t1,...,tn) to (f, [t1,...,tn]) ; naturally tail-recursive*) @@ -412,7 +412,7 @@ | inst (Abs(a,t)) = inst t | inst (f $ u) = (inst f; inst u) | inst _ = () - fun revert() = seq (fn v => v:=NONE) (!updated) + fun revert() = List.app (fn v => v:=NONE) (!updated) in inst t; revert end; @@ -635,7 +635,7 @@ | printPairs _ = () fun printBrs (brs0 as {pairs, lits, lim, ...} :: brs) = (fullTrace := brs0 :: !fullTrace; - seq (fn _ => immediate_output "+") brs; + List.app (fn _ => immediate_output "+") brs; immediate_output (" [" ^ Int.toString lim ^ "] "); printPairs pairs; writeln"") @@ -652,7 +652,7 @@ | 1 => immediate_output"\t1 variable UPDATED:" | n => immediate_output("\t" ^ Int.toString n ^ " variables UPDATED:"); (*display the instantiations themselves, though no variable names*) - seq (fn v => immediate_output(" " ^ string_of sign 4 (the (!v)))) + List.app (fn v => immediate_output(" " ^ string_of sign 4 (valOf (!v)))) (List.take(!trail, !ntrail-ntrl)); writeln"") else (); @@ -751,8 +751,8 @@ end (*substitute throughout "stack frame"; extract affected formulae*) fun subFrame ((Gs,Hs), (changed, frames)) = - let val (changed', Gs') = foldr subForm (Gs, (changed, [])) - val (changed'', Hs') = foldr subForm (Hs, (changed', [])) + let val (changed', Gs') = Library.foldr subForm (Gs, (changed, [])) + val (changed'', Hs') = Library.foldr subForm (Hs, (changed', [])) in (changed'', (Gs',Hs')::frames) end (*substitute throughout literals; extract affected ones*) fun subLit (lit, (changed, nlits)) = @@ -760,8 +760,8 @@ in if nlit aconv lit then (changed, nlit::nlits) else ((nlit,true)::changed, nlits) end - val (changed, lits') = foldr subLit (lits, ([], [])) - val (changed', pairs') = foldr subFrame (pairs, (changed, [])) + val (changed, lits') = Library.foldr subLit (lits, ([], [])) + val (changed', pairs') = Library.foldr subFrame (pairs, (changed, [])) in if !trace then writeln ("Substituting " ^ traceTerm sign u ^ " for " ^ traceTerm sign t ^ " in branch" ) else (); @@ -974,7 +974,7 @@ then lim - (1+log(length rules)) else lim (*discourage branching updates*) val vars = vars_in_vars vars - val vars' = foldr add_terms_vars (prems, vars) + val vars' = Library.foldr add_terms_vars (prems, vars) val choices' = (ntrl, nbrs, PRV) :: choices val tacs' = (tac(updated,false,true)) :: tacs (*no duplication; rotate*) @@ -1101,7 +1101,7 @@ then let val updated = ntrl < !ntrail (*branch updated*) val vars = vars_in_vars vars - val vars' = foldr add_terms_vars (prems, vars) + val vars' = Library.foldr add_terms_vars (prems, vars) (*duplicate H if md permits*) val dup = md (*earlier had "andalso vars' <> vars": duplicate only if the subgoal has new vars*) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/clasimp.ML --- a/src/Provers/clasimp.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/clasimp.ML Thu Mar 03 12:43:01 2005 +0100 @@ -151,30 +151,30 @@ in val op addIffs = - foldl (addIff (Classical.addSDs, Classical.addSEs, Classical.addSIs) + Library.foldl (addIff (Classical.addSDs, Classical.addSEs, Classical.addSIs) (Classical.addDs, Classical.addEs, Classical.addIs) Simplifier.addsimps); -val op delIffs = foldl (delIff Classical.delrules Simplifier.delsimps); +val op delIffs = Library.foldl (delIff Classical.delrules Simplifier.delsimps); fun AddIffs thms = store_clasimp (clasimpset () addIffs thms); fun DelIffs thms = store_clasimp (clasimpset () delIffs thms); fun addIffs_global (thy, ths) = - foldl (addIff + Library.foldl (addIff (ContextRules.addXDs_global, ContextRules.addXEs_global, ContextRules.addXIs_global) (ContextRules.addXDs_global, ContextRules.addXEs_global, ContextRules.addXIs_global) #1) ((thy, ()), ths) |> #1; fun addIffs_local (ctxt, ths) = - foldl (addIff + Library.foldl (addIff (ContextRules.addXDs_local, ContextRules.addXEs_local, ContextRules.addXIs_local) (ContextRules.addXDs_local, ContextRules.addXEs_local, ContextRules.addXIs_local) #1) ((ctxt, ()), ths) |> #1; fun delIffs_global (thy, ths) = - foldl (delIff ContextRules.delrules_global #1) ((thy, ()), ths) |> #1; + Library.foldl (delIff ContextRules.delrules_global #1) ((thy, ()), ths) |> #1; fun delIffs_local (ctxt, ths) = - foldl (delIff ContextRules.delrules_local #1) ((ctxt, ()), ths) |> #1; + Library.foldl (delIff ContextRules.delrules_local #1) ((ctxt, ()), ths) |> #1; end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/classical.ML --- a/src/Provers/classical.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/classical.ML Thu Mar 03 12:43:01 2005 +0100 @@ -214,7 +214,7 @@ let fun addrl (rl,brls) = (false, rl) :: (true, rl RSN (2,swap)) :: brls in assume_tac ORELSE' contr_tac ORELSE' - biresolve_tac (foldr addrl (rls,[])) + biresolve_tac (Library.foldr addrl (rls,[])) end; (*Duplication of hazardous rules, for complete provers*) @@ -286,7 +286,7 @@ fun rep_cs (CS args) = args; local - fun wrap l tac = foldr (fn ((name,tacf),w) => tacf w) (l, tac); + fun wrap l tac = Library.foldr (fn ((name,tacf),w) => tacf w) (l, tac); in fun appSWrappers (CS{swrappers,...}) = wrap swrappers; fun appWrappers (CS{uwrappers,...}) = wrap uwrappers; @@ -316,7 +316,7 @@ fun tag_brls' _ _ [] = [] | tag_brls' w k (brl::brls) = ((w, k), brl) :: tag_brls' w (k + 1) brls; -fun insert_tagged_list kbrls netpr = foldr Tactic.insert_tagged_brl (kbrls, netpr); +fun insert_tagged_list kbrls netpr = Library.foldr Tactic.insert_tagged_brl (kbrls, netpr); (*Insert into netpair that already has nI intr rules and nE elim rules. Count the intr rules double (to account for swapify). Negate to give the @@ -324,7 +324,7 @@ fun insert (nI, nE) = insert_tagged_list o (tag_brls (~(2*nI+nE))) o joinrules; fun insert' w (nI, nE) = insert_tagged_list o tag_brls' w (~(nI + nE)) o joinrules'; -fun delete_tagged_list brls netpr = foldr Tactic.delete_tagged_brl (brls, netpr); +fun delete_tagged_list brls netpr = Library.foldr Tactic.delete_tagged_brl (brls, netpr); fun delete x = delete_tagged_list (joinrules x); fun delete' x = delete_tagged_list (joinrules' x); @@ -355,7 +355,7 @@ cs) else let val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*) - partition Thm.no_prems [th] + List.partition Thm.no_prems [th] val nI = length safeIs + 1 and nE = length safeEs in warn_dup th cs; @@ -380,7 +380,7 @@ cs) else let val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*) - partition (fn rl => nprems_of rl=1) [th] + List.partition (fn rl => nprems_of rl=1) [th] val nI = length safeIs and nE = length safeEs + 1 in warn_dup th cs; @@ -397,7 +397,7 @@ xtra_netpair = insert' 0 (nI,nE) ([], [th]) xtra_netpair} end; -fun rev_foldl f (e, l) = foldl f (e, rev l); +fun rev_foldl f (e, l) = Library.foldl f (e, rev l); val op addSIs = rev_foldl addSI; val op addSEs = rev_foldl addSE; @@ -470,7 +470,7 @@ (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) = if mem_thm (th, safeIs) then - let val (safe0_rls, safep_rls) = partition Thm.no_prems [th] + let val (safe0_rls, safep_rls) = List.partition Thm.no_prems [th] in CS{safe0_netpair = delete (safe0_rls, []) safe0_netpair, safep_netpair = delete (safep_rls, []) safep_netpair, safeIs = rem_thm (safeIs,th), @@ -489,7 +489,7 @@ (cs as CS{safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) = if mem_thm (th, safeEs) then - let val (safe0_rls, safep_rls) = partition (fn rl => nprems_of rl=1) [th] + let val (safe0_rls, safep_rls) = List.partition (fn rl => nprems_of rl=1) [th] in CS{safe0_netpair = delete ([], safe0_rls) safe0_netpair, safep_netpair = delete ([], safep_rls) safep_netpair, safeIs = safeIs, @@ -550,7 +550,7 @@ else (warning ("Undeclared classical rule\n" ^ (string_of_thm th)); cs) end; -val op delrules = foldl delrule; +val op delrules = Library.foldl delrule; (*** Modifying the wrapper tacticals ***) @@ -661,7 +661,7 @@ (*version of bimatch_from_nets_tac that only applies rules that create precisely n subgoals.*) fun n_bimatch_from_nets_tac n = - biresolution_from_nets_tac (Tactic.orderlist o filter (nsubgoalsP n)) true; + biresolution_from_nets_tac (Tactic.orderlist o List.filter (nsubgoalsP n)) true; fun eq_contr_tac i = ematch_tac [not_elim] i THEN eq_assume_tac i; val eq_assume_contr_tac = eq_assume_tac ORELSE' eq_contr_tac; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/ind.ML --- a/src/Provers/ind.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/ind.ML Thu Mar 03 12:43:01 2005 +0100 @@ -41,9 +41,9 @@ (*Generalizes over all free variables, with the named var outermost.*) fun all_frees_tac (var:string) i thm = let val tsig = Sign.tsig_of (Thm.sign_of_thm thm); - val frees = add_term_frees tsig (nth_elem(i-1,prems_of thm),[var]); + val frees = add_term_frees tsig (List.nth(prems_of thm,i-1),[var]); val frees' = sort (rev_order o string_ord) (frees \ var) @ [var] - in foldl (qnt_tac i) (all_tac,frees') thm end; + in Library.foldl (qnt_tac i) (all_tac,frees') thm end; fun REPEAT_SIMP_TAC simp_tac n i = let fun repeat thm = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/induct_method.ML --- a/src/Provers/induct_method.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/induct_method.ML Thu Mar 03 12:43:01 2005 +0100 @@ -59,7 +59,7 @@ val xs = InductAttrib.vars_of tm; in align "Rule has fewer variables than instantiations given" xs ts - |> mapfilter prep_var + |> List.mapPartial prep_var end; @@ -96,7 +96,7 @@ fun inst_rule r = if null insts then RuleCases.add r else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts - |> (flat o map (prep_inst align_left cert I)) + |> (List.concat o map (prep_inst align_left cert I)) |> Drule.cterm_instantiate) r |> rpair (RuleCases.get r); val ruleq = @@ -108,8 +108,8 @@ end | SOME r => Seq.single (inst_rule r)); - fun prep_rule (th, (cases, n)) = Seq.map (apsnd (rpair (drop (n, facts))) o rpair cases) - (Method.multi_resolves (take (n, facts)) [th]); + fun prep_rule (th, (cases, n)) = Seq.map (apsnd (rpair (Library.drop (n, facts))) o rpair cases) + (Method.multi_resolves (Library.take (n, facts)) [th]); in resolveq_cases_tac (RuleCases.make is_open NONE) (Seq.flat (Seq.map prep_rule ruleq)) end; in @@ -161,7 +161,7 @@ let val th = Thm.permute_prems (i - 1) 1 raw_th; val cprems = Drule.cprems_of th; - val As = take (length cprems - 1, cprems); + val As = Library.take (length cprems - 1, cprems); val C = Thm.cterm_of (Thm.sign_of_thm th) (Var (("C", #maxidx (Thm.rep_thm th) + 1), propT)); val dummy_st = Drule.mk_triv_goal (Drule.list_implies (As, C)); in th COMP Thm.lift_rule (dummy_st, 1) Data.local_impI end; @@ -256,7 +256,7 @@ | rename_params ((y,_)::ys) = y::rename_params ys; fun rename_asm (A:term):term = let val xs = rename_params (Logic.strip_params A) - val xs' = case filter (equal x) xs of + val xs' = case List.filter (equal x) xs of [] => xs | [_] => xs | _ => index 1 xs in Logic.list_rename_params (xs',A) end; fun rename_prop (p:term) = @@ -268,9 +268,9 @@ | rename _ thm = thm; fun find_inductT ctxt insts = - foldr multiply (insts |> mapfilter (fn [] => NONE | ts => last_elem ts) + Library.foldr multiply (insts |> List.mapPartial (fn [] => NONE | ts => List.last ts) |> map (InductAttrib.find_inductT ctxt o fastype_of), [[]]) - |> map join_rules |> flat |> map (rename insts); + |> map join_rules |> List.concat |> map (rename insts); fun find_inductS ctxt (fact :: _) = InductAttrib.find_inductS ctxt fact | find_inductS _ _ = []; @@ -291,7 +291,7 @@ if null insts then r else (align_right "Rule has fewer conclusions than arguments given" (Data.dest_concls (Thm.concl_of r)) insts - |> (flat o map (prep_inst align_right cert (atomize_term sg))) + |> (List.concat o map (prep_inst align_right cert (atomize_term sg))) |> Drule.cterm_instantiate) r); val ruleq = @@ -305,8 +305,8 @@ | SOME r => r |> Seq.THEN (rule_versions, Seq.single o inst_rule)); fun prep_rule (th, (cases, n)) = - Seq.map (rpair (cases, n - length facts, drop (n, facts))) - (Method.multi_resolves (take (n, facts)) [th]); + Seq.map (rpair (cases, n - length facts, Library.drop (n, facts))) + (Method.multi_resolves (Library.take (n, facts)) [th]); val tac = resolveq_cases_tac' RuleCases.make is_open (Seq.flat (Seq.map prep_rule ruleq)); in tac THEN_ALL_NEW_CASES rulify_tac end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/make_elim.ML --- a/src/Provers/make_elim.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/make_elim.ML Thu Mar 03 12:43:01 2005 +0100 @@ -41,7 +41,7 @@ | rl'::_ => making (i+1,rl') in making (2, hd (compose_extra 1 (rl, 1, revcut_rl'))) end handle (*in default, use the previous version, which never fails*) - THM _ => Tactic.make_elim rl | LIST _ => Tactic.make_elim rl; + THM _ => Tactic.make_elim rl | Empty => Tactic.make_elim rl; end diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/order.ML --- a/src/Provers/order.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/order.ML Thu Mar 03 12:43:01 2005 +0100 @@ -91,7 +91,7 @@ (* Internal exception, raised if contradiction ( x < x ) was derived *) fun prove asms = - let fun pr (Asm i) = nth_elem (i, asms) + let fun pr (Asm i) = List.nth (asms, i) | pr (Thm (prfs, thm)) = (map pr prfs) MRS thm in pr end; @@ -437,7 +437,7 @@ (* Compute, for each adjacency list, the list with reversed edges, and concatenate these lists. *) - val flipped = foldr (op @) ((map flip g),nil) + val flipped = Library.foldr (op @) ((map flip g),nil) in assemble g flipped end @@ -475,7 +475,7 @@ let val _ = visited := u :: !visited val descendents = - foldr (fn ((v,l),ds) => if been_visited v then ds + Library.foldr (fn ((v,l),ds) => if been_visited v then ds else v :: dfs_visit g v @ ds) ((adjacent (op aconv) g u) ,nil) in @@ -499,7 +499,7 @@ dfs_visit along with u form a connected component. We collect all the connected components together in a list, which is what is returned. *) - foldl (fn (comps,u) => + Library.foldl (fn (comps,u) => if been_visited u then comps else ((u :: dfs_visit (transpose (op aconv) G) u) :: comps)) (nil,(!finish)) end; @@ -525,7 +525,7 @@ let val _ = visited := u :: !visited val descendents = - foldr (fn ((v,l),ds) => if been_visited v then ds + Library.foldr (fn ((v,l),ds) => if been_visited v then ds else v :: dfs_visit g v @ ds) ( ((adjacent (op =) g u)) ,nil) in descendents end @@ -537,7 +537,7 @@ ListPair.map (fn (a,b) => (a,b)) (0 upto (length components -1), components); fun indexNodes IndexComp = - flat (map (fn (index, comp) => (map (fn v => (v,index)) comp)) IndexComp); + List.concat (map (fn (index, comp) => (map (fn v => (v,index)) comp)) IndexComp); fun getIndex v [] = ~1 | getIndex v ((v',k)::vs) = if v aconv v' then k else getIndex v vs; @@ -651,15 +651,15 @@ val v = upper edge in if (getIndex u ntc) = xi then - (completeTermPath x u (nth_elem(xi, sccSubgraphs)) )@[edge] - @(completeTermPath v y' (nth_elem((getIndex y' ntc),sccSubgraphs)) ) + (completeTermPath x u (List.nth(sccSubgraphs, xi)) )@[edge] + @(completeTermPath v y' (List.nth(sccSubgraphs, getIndex y' ntc)) ) else (rev_completeComponentPath u)@[edge] - @(completeTermPath v y' (nth_elem((getIndex y' ntc),sccSubgraphs)) ) + @(completeTermPath v y' (List.nth(sccSubgraphs, getIndex y' ntc)) ) end in if x aconv y then [(Le (x, y, (Thm ([], Less.le_refl))))] - else ( if xi = yi then completeTermPath x y (nth_elem(xi, sccSubgraphs)) + else ( if xi = yi then completeTermPath x y (List.nth(sccSubgraphs, xi)) else rev_completeComponentPath y ) end; @@ -820,15 +820,15 @@ (* if SCC_x is linked to SCC_y via edge e *) else if ui = xi andalso vi = yi then ( case e of (Less (_, _,_)) => ( - let val xypath = (completeTermPath x u (nth_elem(ui, sccSubgraphs)) ) @ [e] @ (completeTermPath v y (nth_elem(vi, sccSubgraphs)) ) + let val xypath = (completeTermPath x u (List.nth(sccSubgraphs, ui)) ) @ [e] @ (completeTermPath v y (List.nth(sccSubgraphs, vi)) ) val xyLesss = transPath (tl xypath, hd xypath) in (Thm ([getprf xyLesss], Less.less_imp_neq)) end) | _ => ( let - val xupath = completeTermPath x u (nth_elem(ui, sccSubgraphs)) - val uxpath = completeTermPath u x (nth_elem(ui, sccSubgraphs)) - val vypath = completeTermPath v y (nth_elem(vi, sccSubgraphs)) - val yvpath = completeTermPath y v (nth_elem(vi, sccSubgraphs)) + val xupath = completeTermPath x u (List.nth(sccSubgraphs, ui)) + val uxpath = completeTermPath u x (List.nth(sccSubgraphs, ui)) + val vypath = completeTermPath v y (List.nth(sccSubgraphs, vi)) + val yvpath = completeTermPath y v (List.nth(sccSubgraphs, vi)) val xuLesss = transPath (tl xupath, hd xupath) val uxLesss = transPath (tl uxpath, hd uxpath) val vyLesss = transPath (tl vypath, hd vypath) @@ -841,15 +841,15 @@ ) ) else if ui = yi andalso vi = xi then ( case e of (Less (_, _,_)) => ( - let val xypath = (completeTermPath y u (nth_elem(ui, sccSubgraphs)) ) @ [e] @ (completeTermPath v x (nth_elem(vi, sccSubgraphs)) ) + let val xypath = (completeTermPath y u (List.nth(sccSubgraphs, ui)) ) @ [e] @ (completeTermPath v x (List.nth(sccSubgraphs, vi)) ) val xyLesss = transPath (tl xypath, hd xypath) in (Thm ([(Thm ([getprf xyLesss], Less.less_imp_neq))] , thm "not_sym")) end ) | _ => ( - let val yupath = completeTermPath y u (nth_elem(ui, sccSubgraphs)) - val uypath = completeTermPath u y (nth_elem(ui, sccSubgraphs)) - val vxpath = completeTermPath v x (nth_elem(vi, sccSubgraphs)) - val xvpath = completeTermPath x v (nth_elem(vi, sccSubgraphs)) + let val yupath = completeTermPath y u (List.nth(sccSubgraphs, ui)) + val uypath = completeTermPath u y (List.nth(sccSubgraphs, ui)) + val vxpath = completeTermPath v x (List.nth(sccSubgraphs, vi)) + val xvpath = completeTermPath x v (List.nth(sccSubgraphs, vi)) val yuLesss = transPath (tl yupath, hd yupath) val uyLesss = transPath (tl uypath, hd uypath) val vxLesss = transPath (tl vxpath, hd vxpath) @@ -944,9 +944,9 @@ fun processComponent (k, comp) = let (* all edges with weight <= of the actual component *) - val leqEdges = flat (map (adjacent (op aconv) leqG) comp); + val leqEdges = List.concat (map (adjacent (op aconv) leqG) comp); (* all edges with weight ~= of the actual component *) - val neqEdges = map snd (flat (map (adjacent (op aconv) neqG) comp)); + val neqEdges = map snd (List.concat (map (adjacent (op aconv) neqG) comp)); (* find an edge leading to a contradiction *) fun findContr [] = NONE @@ -1064,7 +1064,7 @@ val rfrees = map Free (rename_wrt_term A (Logic.strip_params A)) val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A) val C = subst_bounds (rfrees, Logic.strip_assums_concl A) - val lesss = flat (ListPair.map (mkasm_partial sign) (Hs, 0 upto (length Hs - 1))) + val lesss = List.concat (ListPair.map (mkasm_partial sign) (Hs, 0 upto (length Hs - 1))) val prfs = solvePartialOrder sign (lesss, C); val (subgoals, prf) = mkconcl_partial sign C; in @@ -1082,7 +1082,7 @@ val rfrees = map Free (rename_wrt_term A (Logic.strip_params A)) val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A) val C = subst_bounds (rfrees, Logic.strip_assums_concl A) - val lesss = flat (ListPair.map (mkasm_linear sign) (Hs, 0 upto (length Hs - 1))) + val lesss = List.concat (ListPair.map (mkasm_linear sign) (Hs, 0 upto (length Hs - 1))) val prfs = solveTotalOrder sign (lesss, C); val (subgoals, prf) = mkconcl_linear sign C; in diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/quasi.ML --- a/src/Provers/quasi.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/quasi.ML Thu Mar 03 12:43:01 2005 +0100 @@ -92,7 +92,7 @@ (* Internal exception, raised if contradiction ( x < x ) was derived *) fun prove asms = - let fun pr (Asm i) = nth_elem (i, asms) + let fun pr (Asm i) = List.nth (asms, i) | pr (Thm (prfs, thm)) = (map pr prfs) MRS thm in pr end; @@ -560,7 +560,7 @@ val rfrees = map Free (rename_wrt_term A (Logic.strip_params A)) val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A) val C = subst_bounds (rfrees, Logic.strip_assums_concl A) - val lesss = flat (ListPair.map (mkasm_trans sign) (Hs, 0 upto (length Hs - 1))) + val lesss = List.concat (ListPair.map (mkasm_trans sign) (Hs, 0 upto (length Hs - 1))) val prfs = solveLeTrans sign (lesss, C); val (subgoal, prf) = mkconcl_trans sign C; @@ -581,7 +581,7 @@ val rfrees = map Free (rename_wrt_term A (Logic.strip_params A)) val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A) val C = subst_bounds (rfrees, Logic.strip_assums_concl A) - val lesss = flat (ListPair.map (mkasm_quasi sign) (Hs, 0 upto (length Hs - 1))) + val lesss = List.concat (ListPair.map (mkasm_quasi sign) (Hs, 0 upto (length Hs - 1))) val prfs = solveQuasiOrder sign (lesss, C); val (subgoals, prf) = mkconcl_quasi sign C; in diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/simp.ML --- a/src/Provers/simp.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/simp.ML Thu Mar 03 12:43:01 2005 +0100 @@ -78,7 +78,7 @@ bimatch_tac (Net.match_term net (lhs_of (strip_assums_concl prem))) i); -fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); +fun nth_subgoal i thm = List.nth(prems_of thm,i-1); fun goal_concl i thm = strip_assums_concl(nth_subgoal i thm); @@ -155,7 +155,7 @@ in case f of Const(c,T) => if c mem ccs - then foldr add_hvars (args,hvars) + then Library.foldr add_hvars (args,hvars) else add_term_vars(tm,hvars) | _ => add_term_vars(tm,hvars) end @@ -167,7 +167,7 @@ if at then vars else add_term_vars(tm,vars) fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm in if length(tml)=length(al) - then foldr itf (tml~~al,vars) + then Library.foldr itf (tml~~al,vars) else vars end fun add_vars (tm,vars) = case tm of @@ -188,12 +188,12 @@ val lhs = rhs_of_eq 1 thm' val rhs = lhs_of_eq nops thm' val asms = tl(rev(tl(prems_of thm'))) - val hvars = foldr (add_hidden_vars ccs) (lhs::rhs::asms,[]) + val hvars = Library.foldr (add_hidden_vars ccs) (lhs::rhs::asms,[]) val hvars = add_new_asm_vars new_asms (rhs,hvars) fun it_asms (asm,hvars) = if atomic asm then add_new_asm_vars new_asms (asm,hvars) else add_term_frees(asm,hvars) - val hvars = foldr it_asms (asms,hvars) + val hvars = Library.foldr it_asms (asms,hvars) val hvs = map (#1 o dest_Var) hvars val refl1_tac = refl_tac 1 fun norm_step_tac st = st |> @@ -210,7 +210,7 @@ fun add_norm_tags congs = let val ccs = map cong_const congs - val new_asms = filter (exists not o #2) + val new_asms = List.filter (exists not o #2) (ccs ~~ (map (map atomic o prems_of) congs)); in add_norms(congs,ccs,new_asms) end; @@ -249,7 +249,7 @@ (writeln"\nDuplicate rewrite or congruence rule:"; print_thm th; net); -val insert_thms = foldr insert_thm_warn; +val insert_thms = Library.foldr insert_thm_warn; fun addrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, splits,split_consts}, thm) = @@ -259,7 +259,7 @@ splits=splits,split_consts=split_consts} end; -val op addrews = foldl addrew; +val op addrews = Library.foldl addrew; fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, splits,split_consts}, thms) = @@ -283,7 +283,7 @@ mk_simps=mk_simps,simps=simps,simp_net=simp_net, splits=splits@[mk_trans thm],split_consts=split_consts@[a]} end; -val op addsplits = foldl addsplit; +val op addsplits = Library.foldl addsplit; (** Deletion of congruences and rewrites **) @@ -294,11 +294,11 @@ (writeln"\nNo such rewrite or congruence rule:"; print_thm th; net); -val delete_thms = foldr delete_thm_warn; +val delete_thms = Library.foldr delete_thm_warn; fun op delcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net, splits,split_consts}, thms) = -let val congs' = foldl (gen_rem Drule.eq_thm_prop) (congs,thms) +let val congs' = Library.foldl (gen_rem Drule.eq_thm_prop) (congs,thms) in SS{auto_tac=auto_tac, congs= congs', cong_net= delete_thms(map mk_trans thms,cong_net), mk_simps= normed_rews congs', simps=simps, simp_net=simp_net, @@ -318,7 +318,7 @@ splits=splits,split_consts=split_consts} end; -val op delrews = foldl delrew; +val op delrews = Library.foldl delrew; fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net, @@ -447,9 +447,9 @@ SOME(thm',_) => let val ps = prems_of thm and ps' = prems_of thm'; val n = length(ps')-length(ps); - val a = length(strip_assums_hyp(nth_elem(i-1,ps))) + val a = length(strip_assums_hyp(List.nth(ps,i-1))) val l = map (fn p => length(strip_assums_hyp(p))) - (take(n,drop(i-1,ps'))); + (Library.take(n,Library.drop(i-1,ps'))); in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end | NONE => (REW::ss,thm,anet,ats,cs); @@ -460,7 +460,7 @@ val thms = map (trivial o cterm_of(#sign(rep_thm(thm))))As; val new_rws = flat(map mk_rew_rules thms); val rwrls = map mk_trans (flat(map mk_rew_rules thms)); - val anet' = foldr lhs_insert_thm (rwrls,anet) + val anet' = Library.foldr lhs_insert_thm (rwrls,anet) in if !tracing andalso not(null new_rws) then (writeln"Adding rewrites:"; prths new_rws; ()) else (); @@ -565,7 +565,7 @@ fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) = let fun xn_list(x,n) = let val ixs = map (fn i => (x^(radixstring(26,"a",i)),0)) (0 upto n); - in ListPair.map eta_Var (ixs, take(n+1,Ts)) end + in ListPair.map eta_Var (ixs, Library.take(n+1,Ts)) end val lhs = list_comb(f,xn_list("X",k-1)) val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik) in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/splitter.ML --- a/src/Provers/splitter.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/splitter.ML Thu Mar 03 12:43:01 2005 +0100 @@ -102,8 +102,8 @@ fun down [] t i = Bound 0 | down (p::ps) t i = let val (h,ts) = strip_comb t - val v1 = ListPair.map var (take(p,ts), i upto (i+p-1)) - val u::us = drop(p,ts) + val v1 = ListPair.map var (Library.take(p,ts), i upto (i+p-1)) + val u::us = Library.drop(p,ts) val v2 = ListPair.map var (us, (i+p) upto (i+length(ts)-2)) in list_comb(h,v1@[down ps u (i+length ts)]@v2) end; in Abs("", T, down (rev pos) t maxi) end; @@ -137,7 +137,7 @@ has a body of type T; otherwise the expansion thm will fail later on *) fun type_test(T,lbnos,apsns) = - let val (_,U,_) = nth_elem(foldl Int.min (hd lbnos, tl lbnos), apsns) + let val (_,U,_) = List.nth(apsns, Library.foldl Int.min (hd lbnos, tl lbnos)) in T=U end; (************************************************************************* @@ -171,8 +171,8 @@ fun mk_split_pack(thm, T, T', n, ts, apsns, pos, TB, t) = if n > length ts then [] else let val lev = length apsns - val lbnos = foldl add_lbnos ([],take(n,ts)) - val flbnos = filter (fn i => i < lev) lbnos + val lbnos = Library.foldl add_lbnos ([],Library.take(n,ts)) + val flbnos = List.filter (fn i => i < lev) lbnos val tt = incr_boundvars (~lev) t in if null flbnos then if T = T' then [(thm,[],pos,TB,tt)] else [] @@ -234,7 +234,7 @@ Const(c, cT) => let fun find [] = [] | find ((gcT, pat, thm, T, n)::tups) = - let val t2 = list_comb (h, take (n, ts)) + let val t2 = list_comb (h, Library.take (n, ts)) in if Sign.typ_instance sg (cT, gcT) andalso fomatch (Sign.tsig_of sg) (Ts,pat,t2) then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2) @@ -242,11 +242,11 @@ end in find (assocs cmap c) end | _ => [] - in snd(foldl iter ((0, a), ts)) end + in snd(Library.foldl iter ((0, a), ts)) end in posns Ts [] [] t end; -fun nth_subgoal i thm = nth_elem(i-1,prems_of thm); +fun nth_subgoal i thm = List.nth(prems_of thm,i-1); fun shorter((_,ps,pos,_,_),(_,qs,qos,_,_)) = prod_ord (int_ord o pairself length) (order o pairself length) @@ -312,7 +312,7 @@ (Logic.strip_assums_concl (#prop (rep_thm thm'))))); val cert = cterm_of (sign_of_thm state); val cntxt = mk_cntxt_splitthm t tt TB; - val abss = foldl (fn (t, T) => Abs ("", T, t)); + val abss = Library.foldl (fn (t, T) => Abs ("", T, t)); in cterm_instantiate [(cert P, cert (abss (cntxt, Ts)))] thm' end; @@ -337,7 +337,7 @@ end | _ => split_format_err()) | _ => split_format_err()) - val cmap = foldl add_thm ([],splits); + val cmap = Library.foldl add_thm ([],splits); fun lift_tac Ts t p st = rtac (inst_lift Ts t p st i) i st fun lift_split_tac state = let val (Ts, t, splits) = select cmap state i @@ -421,13 +421,13 @@ let val (name,asm) = split_thm_info split in Simplifier.addloop (ss, (split_name name asm, (if asm then split_asm_tac else split_tac) [split])) end - in foldl addsplit (ss,splits) end; + in Library.foldl addsplit (ss,splits) end; fun ss delsplits splits = let fun delsplit(ss,split) = let val (name,asm) = split_thm_info split in Simplifier.delloop (ss, split_name name asm) - end in foldl delsplit (ss,splits) end; + end in Library.foldl delsplit (ss,splits) end; fun Addsplits splits = (Simplifier.simpset_ref() := Simplifier.simpset() addsplits splits); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/trancl.ML --- a/src/Provers/trancl.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/trancl.ML Thu Mar 03 12:43:01 2005 +0100 @@ -88,7 +88,7 @@ exception Cannot; (* internal exception: raised if no proof can be found *) fun prove asms = - let fun pr (Asm i) = nth_elem (i, asms) + let fun pr (Asm i) = List.nth (asms, i) | pr (Thm (prfs, thm)) = (map pr prfs) MRS thm in pr end; @@ -327,7 +327,7 @@ (* Compute, for each adjacency list, the list with reversed edges, and concatenate these lists. *) - val flipped = foldr (op @) ((map flip g),nil) + val flipped = Library.foldr (op @) ((map flip g),nil) in assemble g flipped end @@ -351,7 +351,7 @@ let val _ = visited := u :: !visited val descendents = - foldr (fn ((v,l),ds) => if been_visited v then ds + Library.foldr (fn ((v,l),ds) => if been_visited v then ds else v :: dfs_visit g v @ ds) ( ((adjacent eq_comp g u)) ,nil) in descendents end @@ -530,7 +530,7 @@ val Hs = Logic.strip_assums_hyp A; val C = Logic.strip_assums_concl A; val (rel,subgoals, prf) = mkconcl_trancl C; - val prems = flat (ListPair.map (mkasm_trancl rel) (Hs, 0 upto (length Hs - 1))) + val prems = List.concat (ListPair.map (mkasm_trancl rel) (Hs, 0 upto (length Hs - 1))) val prfs = solveTrancl (prems, C); in @@ -548,7 +548,7 @@ val C = Logic.strip_assums_concl A; val (rel,subgoals, prf) = mkconcl_rtrancl C; - val prems = flat (ListPair.map (mkasm_rtrancl rel) (Hs, 0 upto (length Hs - 1))) + val prems = List.concat (ListPair.map (mkasm_rtrancl rel) (Hs, 0 upto (length Hs - 1))) val prfs = solveRtrancl (prems, C); in METAHYPS (fn asms => diff -r 1b3115d1a8df -r 8d8c70b41bab src/Provers/typedsimp.ML --- a/src/Provers/typedsimp.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Provers/typedsimp.ML Thu Mar 03 12:43:01 2005 +0100 @@ -70,7 +70,7 @@ handle THM _ => (simp_rls, rl :: other_rls); (*Given the list rls, return the pair (simp_rls, other_rls).*) -fun process_rules rls = foldr add_rule (rls, ([],[])); +fun process_rules rls = Library.foldr add_rule (rls, ([],[])); (*Given list of rewrite rules, return list of both forms, reject others*) fun process_rewrites rls = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/file.ML --- a/src/Pure/General/file.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/file.ML Thu Mar 03 12:43:01 2005 +0100 @@ -96,7 +96,7 @@ | s => SOME (Info s)) end; -val exists = is_some o info; +val exists = isSome o info; (* directories *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/graph.ML --- a/src/Pure/General/graph.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/graph.ML Thu Mar 03 12:43:01 2005 +0100 @@ -67,7 +67,7 @@ val empty_keys = Table.empty: keys; infix mem_keys; -fun x mem_keys tab = is_some (Table.lookup (tab: keys, x)); +fun x mem_keys tab = isSome (Table.lookup (tab: keys, x)); infix ins_keys; fun x ins_keys tab = if x mem_keys tab then tab else Table.update ((x, ()), tab); @@ -114,7 +114,7 @@ fun reach ((R, rs), x) = if x mem_keys R then (R, rs) else apsnd (cons x) (reachs ((x ins_keys R, rs), next x)) - and reachs R_xs = foldl reach R_xs; + and reachs R_xs = Library.foldl reach R_xs; in foldl_map (reach o apfst (rpair [])) (empty_keys, xs) end; (*immediate*) @@ -122,13 +122,13 @@ fun imm_succs G = #2 o #2 o get_entry G; (*transitive*) -fun all_preds G = flat o snd o reachable (imm_preds G); -fun all_succs G = flat o snd o reachable (imm_succs G); +fun all_preds G = List.concat o snd o reachable (imm_preds G); +fun all_succs G = List.concat o snd o reachable (imm_succs G); (*strongly connected components; see: David King and John Launchbury, "Structuring Depth First Search Algorithms in Haskell"*) fun strong_conn G = filter_out null (snd (reachable (imm_preds G) - (flat (rev (snd (reachable (imm_succs G) (keys G))))))); + (List.concat (rev (snd (reachable (imm_succs G) (keys G))))))); (* paths *) @@ -139,7 +139,7 @@ fun paths ps p = if not (null ps) andalso eq_key (p, x) then [p :: ps] else if p mem_keys X andalso not (p mem_key ps) - then flat (map (paths (p :: ps)) (imm_preds G p)) + then List.concat (map (paths (p :: ps)) (imm_preds G p)) else []; in paths [] y end; @@ -155,8 +155,8 @@ let fun del (x, (i, (preds, succs))) = if x mem_key xs then NONE - else SOME (x, (i, (foldl op del_key (preds, xs), foldl op del_key (succs, xs)))); - in Graph (Table.make (mapfilter del (Table.dest tab))) end; + else SOME (x, (i, (Library.foldl op del_key (preds, xs), Library.foldl op del_key (succs, xs)))); + in Graph (Table.make (List.mapPartial del (Table.dest tab))) end; (* edges *) @@ -176,7 +176,7 @@ else G; fun diff_edges G1 G2 = - flat (dest G1 |> map (fn (x, ys) => ys |> mapfilter (fn y => + List.concat (dest G1 |> map (fn (x, ys) => ys |> List.mapPartial (fn y => if is_edge G2 (x, y) then NONE else SOME (x, y)))); fun edges G = diff_edges G empty; @@ -205,7 +205,7 @@ | cycles => raise CYCLES (map (cons x) cycles)); fun add_deps_acyclic (y, xs) G = - foldl (fn (H, x) => add_edge_acyclic (x, y) H) (G, xs); + Library.foldl (fn (H, x) => add_edge_acyclic (x, y) H) (G, xs); fun merge_acyclic eq GG = gen_merge add_edge_acyclic eq GG; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/history.ML --- a/src/Pure/General/history.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/history.ML Thu Mar 03 12:43:01 2005 +0100 @@ -35,7 +35,7 @@ fun current (History (x, _)) = x; fun clear n (History (x, (lim, len, undo_list, redo_list))) = - History (x, (lim, Int.max (0, len - n), drop (n, undo_list), redo_list)); + History (x, (lim, Int.max (0, len - n), Library.drop (n, undo_list), redo_list)); fun undo (History (_, (_, _, [], _))) = error "No further undo information" | undo (History (x, (lim, len, u :: undo_list, redo_list))) = @@ -47,7 +47,7 @@ fun push NONE _ x xs = x :: xs | push (SOME 0) _ _ _ = [] - | push (SOME n) len x xs = if len < n then x :: xs else take (n, x :: xs); + | push (SOME n) len x xs = if len < n then x :: xs else Library.take (n, x :: xs); fun apply_copy cp f (History (x, (lim, len, undo_list, _))) = let val x' = cp x diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/lazy_seq.ML --- a/src/Pure/General/lazy_seq.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/lazy_seq.ML Thu Mar 03 12:43:01 2005 +0100 @@ -83,7 +83,7 @@ fun getItem (Seq s) = force s fun make f = Seq (delay f) -fun null s = is_some (getItem s) +fun null s = isSome (getItem s) local fun F n NONE = n @@ -397,8 +397,8 @@ make (fn () => copy (f x)) end -fun EVERY fs = foldr THEN (fs, succeed) -fun FIRST fs = foldr ORELSE (fs, fail) +fun EVERY fs = Library.foldr THEN (fs, succeed) +fun FIRST fs = Library.foldr ORELSE (fs, fail) fun TRY f x = make (fn () => diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/name_space.ML --- a/src/Pure/General/name_space.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/name_space.ML Thu Mar 03 12:43:01 2005 +0100 @@ -50,7 +50,7 @@ val unpack = space_explode separator; val pack = space_implode separator; -val base = last_elem o unpack; +val base = List.last o unpack; fun map_base f name = let val uname = unpack name @@ -75,7 +75,7 @@ (* basic operations *) fun map_space f xname tab = - Symtab.update ((xname, f (if_none (Symtab.lookup (tab, xname)) ([], []))), tab); + Symtab.update ((xname, f (getOpt (Symtab.lookup (tab, xname), ([], [])))), tab); fun change f xname (name, tab) = map_space (fn (names, names') => (f names name, names')) xname tab; @@ -100,9 +100,9 @@ fun extend_aux (name, tab) = if is_hidden name then error ("Attempt to declare hidden name " ^ quote name) - else foldl (fn (tb, xname) => change add xname (name, tb)) (tab, accesses name); + else Library.foldl (fn (tb, xname) => change add xname (name, tb)) (tab, accesses name); -fun extend (NameSpace tab, names) = NameSpace (foldr extend_aux (names, tab)); +fun extend (NameSpace tab, names) = NameSpace (Library.foldr extend_aux (names, tab)); (* merge *) (*1st preferred over 2nd*) @@ -123,10 +123,10 @@ else if is_hidden name then error ("Attempt to hide hidden name " ^ quote name) else (change' add name (name, - foldl (fn (tb, xname) => change del xname (name, tb)) + Library.foldl (fn (tb, xname) => change del xname (name, tb)) (tab, if fully then accesses name else [base name]))); -fun hide fully (NameSpace tab, names) = NameSpace (foldr (hide_aux fully) (names, tab)); +fun hide fully (NameSpace tab, names) = NameSpace (Library.foldr (hide_aux fully) (names, tab)); (* intern / extern names *) @@ -175,7 +175,7 @@ fun dest (NameSpace tab) = map (apsnd (sort (int_ord o pairself size))) - (Symtab.dest (Symtab.make_multi (mapfilter dest_entry (Symtab.dest tab)))); + (Symtab.dest (Symtab.make_multi (List.mapPartial dest_entry (Symtab.dest tab)))); end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/path.ML --- a/src/Pure/General/path.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/path.ML Thu Mar 03 12:43:01 2005 +0100 @@ -90,7 +90,7 @@ | rev_app xs (y :: ys) = rev_app (y :: xs) ys; fun append (Path xs) (Path ys) = Path (rev_app (rev xs) ys); -fun appends paths = foldl (uncurry append) (current, paths); +fun appends paths = Library.foldl (uncurry append) (current, paths); val make = appends o map basic; fun norm path = rev_app [] path; @@ -142,7 +142,7 @@ val split_ext = split_path (fn (prfx, s) => apfst (append (Path prfx)) (case take_suffix (not_equal ".") (explode s) of ([], _) => (Path [Basic s], "") - | (cs, e) => (Path [Basic (implode (take (length cs - 1, cs)))], implode e))); + | (cs, e) => (Path [Basic (implode (Library.take (length cs - 1, cs)))], implode e))); (* evaluate variables *) @@ -151,7 +151,7 @@ | eval _ x = [x]; fun evaluate env (Path xs) = - Path (norm (flat (map (eval env) xs))); + Path (norm (List.concat (map (eval env) xs))); val expand = evaluate (unpack o getenv); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/pretty.ML --- a/src/Pure/General/pretty.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/pretty.ML Thu Mar 03 12:43:01 2005 +0100 @@ -204,7 +204,7 @@ blk (1, [str "\"", prt, str "\""]); fun commas prts = - flat (separate [str ",", brk 1] (map (fn x => [x]) prts)); + List.concat (separate [str ",", brk 1] (map (fn x => [x]) prts)); fun breaks prts = separate (brk 1) prts; fun fbreaks prts = separate fbrk prts; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/scan.ML --- a/src/Pure/General/scan.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/scan.ML Thu Mar 03 12:43:01 2005 +0100 @@ -204,7 +204,7 @@ fun trace scan toks = let val (x, toks') = scan toks - in ((x, take (length toks - length toks', toks)), toks') end; + in ((x, Library.take (length toks - length toks', toks)), toks') end; (* state based scanners *) @@ -247,7 +247,7 @@ in if exists is_stopper input then raise ABORT "Stopper may not occur in input of finite scan!" - else (force scan --| lift stop) (state, rev_append (rev input) [stopper]) + else (force scan --| lift stop) (state, List.revAppend (rev input,[stopper])) end; fun finite stopper scan xs = @@ -264,7 +264,7 @@ fun drain def_prmpt get stopper scan ((state, xs), src) = (scan (state, xs), src) handle MORE prmpt => - (case get (if_none prmpt def_prmpt) src of + (case get (getOpt (prmpt,def_prmpt)) src of ([], _) => (finite' stopper scan (state, xs), src) | (xs', src') => drain def_prmpt get stopper scan ((state, xs @ xs'), src')); @@ -274,7 +274,7 @@ fun drain_loop recover inp = drain_with (catch scanner) inp handle FAIL msg => - (error_msg (if_none msg "Syntax error."); drain_with recover inp); + (error_msg (getOpt (msg,"Syntax error.")); drain_with recover inp); val ((ys, (state', xs')), src') = (case (get def_prmpt src, opt_recover) of @@ -285,7 +285,7 @@ fun source def_prmpt get unget stopper scan opt_recover src = let val (ys, ((), src')) = - source' def_prmpt get unget stopper (lift scan) (apsome lift opt_recover) ((), src) + source' def_prmpt get unget stopper (lift scan) (Option.map lift opt_recover) ((), src) in (ys, src') end; fun single scan = scan >> (fn x => [x]); @@ -333,7 +333,7 @@ Branch (c, no_literal, Empty, add Empty cs, Empty) | add lex [] = lex; in add lex chrs end; - in foldl ext (lexicon, chrss \\ dest_lex lexicon) end; + in Library.foldl ext (lexicon, chrss \\ dest_lex lexicon) end; val make_lexicon = extend_lexicon empty_lexicon; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/seq.ML --- a/src/Pure/General/seq.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/seq.ML Thu Mar 03 12:43:01 2005 +0100 @@ -71,8 +71,8 @@ fun single x = cons (x, empty); (*head and tail -- beware of calling the sequence function twice!!*) -fun hd xq = #1 (the (pull xq)) -and tl xq = #2 (the (pull xq)); +fun hd xq = #1 (valOf (pull xq)) +and tl xq = #2 (valOf (pull xq)); (*partial function as procedure*) fun try f x = @@ -97,7 +97,7 @@ | SOME (x, xq') => x :: list_of xq'); (*conversion from list to sequence*) -fun of_list xs = foldr cons (xs, empty); +fun of_list xs = Library.foldr cons (xs, empty); (*map the function f over the sequence, making a new sequence*) @@ -203,8 +203,8 @@ fun op APPEND (f, g) x = append (f x, make (fn () => pull (g x))); -fun EVERY fs = foldr THEN (fs, succeed); -fun FIRST fs = foldr ORELSE (fs, fail); +fun EVERY fs = Library.foldr THEN (fs, succeed); +fun FIRST fs = Library.foldr ORELSE (fs, fail); fun TRY f = ORELSE (f, succeed); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/symbol.ML --- a/src/Pure/General/symbol.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/symbol.ML Thu Mar 03 12:43:01 2005 +0100 @@ -342,7 +342,7 @@ else if is_ascii_quasi s then Quasi else if is_ascii_blank s then Blank else if is_char s then Other - else if_none (Symtab.lookup (symbol_kinds, s)) Other; + else getOpt (Symtab.lookup (symbol_kinds, s), Other); end; fun is_letter s = kind s = Letter; @@ -417,7 +417,7 @@ fun sym_explode str = let val chs = explode str in if no_explode chs then chs - else the (Scan.read stopper (Scan.repeat scan) chs) + else valOf (Scan.read stopper (Scan.repeat scan) chs) end; val chars_only = not o exists_string (equal "\\"); @@ -486,7 +486,7 @@ else if s = "\\" then 2 else 1; -fun sym_length ss = foldl (fn (n, s) => sym_len s + n) (0, ss); +fun sym_length ss = Library.foldl (fn (n, s) => sym_len s + n) (0, ss); fun symbol_output str = if chars_only str then default_output str diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/table.ML --- a/src/Pure/General/table.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/table.ML Thu Mar 03 12:43:01 2005 +0100 @@ -90,12 +90,12 @@ fun exists P tab = foldl_table (fn (false, e) => P e | (b, _) => b) (false, tab); fun min_key Empty = NONE - | min_key (Branch2 (left, (k, _), _)) = SOME (if_none (min_key left) k) - | min_key (Branch3 (left, (k, _), _, _, _)) = SOME (if_none (min_key left) k); + | min_key (Branch2 (left, (k, _), _)) = SOME (getOpt (min_key left,k)) + | min_key (Branch3 (left, (k, _), _, _, _)) = SOME (getOpt (min_key left,k)); fun max_key Empty = NONE - | max_key (Branch2 (_, (k, _), right)) = SOME (if_none (max_key right) k) - | max_key (Branch3 (_, _, _, (k,_), right)) = SOME (if_none (max_key right) k); + | max_key (Branch2 (_, (k, _), right)) = SOME (getOpt (max_key right,k)) + | max_key (Branch3 (_, _, _, (k,_), right)) = SOME (getOpt (max_key right,k)); (* lookup *) @@ -178,7 +178,7 @@ NONE => (update ((key, x), tab), dups) | _ => (tab, key :: dups)); in - (case foldl add ((table, []), pairs) of + (case Library.foldl add ((table, []), pairs) of (table', []) => table' | (_, dups) => raise DUPS (rev dups)) end; @@ -201,12 +201,12 @@ | del NONE (Branch3 (Empty, p, Empty, q, Empty)) = (p, (false, Branch2 (Empty, q, Empty))) | del k (Branch2 (Empty, p, Empty)) = (case compare' k p of - EQUAL => (p, (true, Empty)) | _ => raise UNDEF (the k)) + EQUAL => (p, (true, Empty)) | _ => raise UNDEF (valOf k)) | del k (Branch3 (Empty, p, Empty, q, Empty)) = (case compare' k p of EQUAL => (p, (false, Branch2 (Empty, q, Empty))) | _ => (case compare' k q of EQUAL => (q, (false, Branch2 (Empty, p, Empty))) - | _ => raise UNDEF (the k))) + | _ => raise UNDEF (valOf k))) | del k (Branch2 (l, p, r)) = (case compare' k p of LESS => (case del k l of (p', (false, l')) => (p', (false, Branch2 (l', p, r))) @@ -284,11 +284,11 @@ (* tables with multiple entries per key (preserving order) *) -fun lookup_multi tab_key = if_none (lookup tab_key) []; +fun lookup_multi tab_key = getOpt (lookup tab_key,[]); fun update_multi ((key, x), tab) = update ((key, x :: lookup_multi (tab, key)), tab); -fun make_multi pairs = foldr update_multi (pairs, empty); -fun dest_multi tab = flat (map (fn (key, xs) => map (Library.pair key) xs) (dest tab)); +fun make_multi pairs = Library.foldr update_multi (pairs, empty); +fun dest_multi tab = List.concat (map (fn (key, xs) => map (Library.pair key) xs) (dest tab)); fun merge_multi eq tabs = join (fn (xs, xs') => SOME (gen_merge_lists eq xs xs')) tabs; fun merge_multi' eq tabs = join (fn (xs, xs') => SOME (gen_merge_lists' eq xs xs')) tabs; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/General/xml.ML --- a/src/Pure/General/xml.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/General/xml.ML Thu Mar 03 12:43:01 2005 +0100 @@ -134,7 +134,7 @@ || parse_pi >> K [] || parse_comment >> K []) -- Scan.optional ((* scan_whspc |-- *) parse_chars >> (single o Text)) [] - >> op @) >> flat) >> op @)(* --| scan_whspc*)) xs + >> op @) >> List.concat) >> op @)(* --| scan_whspc*)) xs and parse_elem xs = ($$ "<" |-- Symbol.scan_id -- diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/IsaPlanner/isand.ML --- a/src/Pure/IsaPlanner/isand.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/IsaPlanner/isand.ML Thu Mar 03 12:43:01 2005 +0100 @@ -108,13 +108,13 @@ fun allify_prem_var (vt as (n,ty),t) = (Term.all ty) $ (Abs(n,ty,Term.abstract_over (Free vt, t))) - fun allify_prem Ts p = foldr allify_prem_var (Ts, p) + fun allify_prem Ts p = Library.foldr allify_prem_var (Ts, p) val cTs = map (ctermify o Free) Ts val cterm_asms = map (ctermify o allify_prem Ts) premts val allifyied_asm_thms = map (Drule.forall_elim_list cTs o Thm.assume) cterm_asms in - (foldl (fn (x,y) => y COMP x) (th, allifyied_asm_thms), cterm_asms) + (Library.foldl (fn (x,y) => y COMP x) (th, allifyied_asm_thms), cterm_asms) end; fun allify_conditions' Ts th = @@ -132,7 +132,7 @@ val names = Term.add_term_names (prop, []) val (insts,names2) = - foldl (fn ((insts,names),v as ((n,i),ty)) => + Library.foldl (fn ((insts,names),v as ((n,i),ty)) => let val n2 = Term.variant names n in ((ctermify (Var v), ctermify (Free(n2,ty))) :: insts, n2 :: names) @@ -167,7 +167,7 @@ in (Term.all ty) $ (Abs(n,ty,Term.abstract_over (vt, t))) end; fun allify_for_sg_term ctermify vs t = - let val t_alls = foldr allify_term (vs,t); + let val t_alls = Library.foldr allify_term (vs,t); val ct_alls = ctermify t_alls; in (ct_alls, Drule.forall_elim_list vs (Thm.assume ct_alls)) @@ -189,7 +189,7 @@ val (sgallcts, sgthms) = Library.split_list (map (allify_for_sg_term ctermify vs) sgs); val minimal_newth = - (foldl (fn ( newth', sgthm) => + (Library.foldl (fn ( newth', sgthm) => Drule.compose_single (sgthm, 1, newth')) (newth, sgthms)); val allified_newth = @@ -229,7 +229,7 @@ Library.split_list (map (allify_for_sg_term ctermify cfrees) sgs); val minimal_th = - (foldl (fn ( th', sgthm) => + (Library.foldl (fn ( th', sgthm) => Drule.compose_single (sgthm, 1, th')) (th, sgthms)) RS Drule.rev_triv_goal; @@ -256,7 +256,7 @@ |> Drule.forall_intr_list cfvs in Drule.compose_single (solth', i, gth) end; -val export_solutions = foldr (uncurry export_solution); +val export_solutions = Library.foldr (uncurry export_solution); (* fix parameters of a subgoal "i", as free variables, and create an diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/IsaPlanner/isaplib.ML --- a/src/Pure/IsaPlanner/isaplib.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/IsaPlanner/isaplib.ML Thu Mar 03 12:43:01 2005 +0100 @@ -159,7 +159,7 @@ (* for use with tactics... gives no_tac if element isn't there *) fun NTH n f st = let val r = nth_of_seq n (f st) in - if (is_none r) then Seq.empty else (Seq.single (the r)) + if (is_none r) then Seq.empty else (Seq.single (valOf r)) end; (* First result of a tactic... uses NTH, so if there is no elements, @@ -214,7 +214,7 @@ else if gr(l,h) then get_ends_of gr (u,h) t else get_ends_of gr (u,l) t; -fun flatmap f = flat o map f; +fun flatmap f = List.concat o map f; (* quick removal of "rs" elements from "ls" when (eqf (r,l)) is true Ignores ordering. *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/IsaPlanner/rw_inst.ML --- a/src/Pure/IsaPlanner/rw_inst.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/IsaPlanner/rw_inst.ML Thu Mar 03 12:43:01 2005 +0100 @@ -89,7 +89,7 @@ val prop = Thm.prop_of rule; val names = Term.add_term_names (prop, []); val (fromnames,tonames,names2,Ts') = - foldl (fn ((rnf,rnt,names, Ts'),(n,ty)) => + Library.foldl (fn ((rnf,rnt,names, Ts'),(n,ty)) => let val n2 = Term.variant names n in (ctermify (Free(RWTools.mk_fake_bound_name n, ty)) :: rnf, @@ -108,7 +108,7 @@ (* using these names create lambda-abstracted version of the rule *) val abstractions = rev (Ts' ~~ tonames); - val abstract_rule = foldl (fn (th,((n,ty),ct)) => + val abstract_rule = Library.foldl (fn (th,((n,ty),ct)) => Thm.abstract_rule n ct th) (uncond_rule, abstractions); in (cprems, abstract_rule) end; @@ -124,9 +124,9 @@ fun mk_renamings tgt rule_inst = let val rule_conds = Thm.prems_of rule_inst - val names = foldr Term.add_term_names (tgt :: rule_conds, []); + val names = Library.foldr Term.add_term_names (tgt :: rule_conds, []); val (conds_tyvs,cond_vs) = - foldl (fn ((tyvs, vs), t) => + Library.foldl (fn ((tyvs, vs), t) => (Library.union (Term.term_tvars t, tyvs), Library.union @@ -135,7 +135,7 @@ val termvars = map Term.dest_Var (Term.term_vars tgt); val vars_to_fix = Library.union (termvars, cond_vs); val (renamings, names2) = - foldr (fn (((n,i),ty), (vs, names')) => + Library.foldr (fn (((n,i),ty), (vs, names')) => let val n' = Term.variant names' n in ((((n,i),ty), Free (n', ty)) :: vs, n'::names') end) @@ -152,12 +152,12 @@ already instantiated (in ignore_ixs) from the list of terms. *) fun mk_fixtvar_tyinsts ignore_ixs ts = let val (tvars, tfreenames) = - foldr (fn (t, (varixs, tfreenames)) => + Library.foldr (fn (t, (varixs, tfreenames)) => (Term.add_term_tvars (t,varixs), Term.add_term_tfree_names (t,tfreenames))) (ts, ([],[])); - val unfixed_tvars = filter (fn (ix,s) => not (ix mem ignore_ixs)) tvars; - val (fixtyinsts, _) = foldr new_tfree (unfixed_tvars, ([], tfreenames)) + val unfixed_tvars = List.filter (fn (ix,s) => not (ix mem ignore_ixs)) tvars; + val (fixtyinsts, _) = Library.foldr new_tfree (unfixed_tvars, ([], tfreenames)) in (fixtyinsts, tfreenames) end; @@ -222,7 +222,7 @@ Term.subst_vars (typinsts,[]) outerterm; (* type-instantiate the var instantiations *) - val insts_tyinst = foldr (fn ((ix,t),insts_tyinst) => + val insts_tyinst = Library.foldr (fn ((ix,t),insts_tyinst) => (ix, Term.subst_vars (typinsts,[]) t) :: insts_tyinst) (unprepinsts,[]); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/IsaPlanner/rw_tools.ML --- a/src/Pure/IsaPlanner/rw_tools.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/IsaPlanner/rw_tools.ML Thu Mar 03 12:43:01 2005 +0100 @@ -130,7 +130,7 @@ fun fake_concl_of_goal gt i = let val prems = Logic.strip_imp_prems gt - val sgt = nth_elem (i - 1, prems) + val sgt = List.nth (prems, i - 1) val tbody = Logic.strip_imp_concl (Term.strip_all_body sgt) val tparams = Term.strip_all_vars sgt @@ -148,7 +148,7 @@ fun fake_goal gt i = let val prems = Logic.strip_imp_prems gt - val sgt = nth_elem (i - 1, prems) + val sgt = List.nth (prems, i - 1) val tbody = Term.strip_all_body sgt val tparams = Term.strip_all_vars sgt diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/IsaPlanner/term_lib.ML --- a/src/Pure/IsaPlanner/term_lib.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/IsaPlanner/term_lib.ML Thu Mar 03 12:43:01 2005 +0100 @@ -163,20 +163,20 @@ case typs_unify of SOME (typinsttab, ix2) => let - (* is it right to throw away teh flexes? + (* is it right to throw away the flexes? or should I be using them somehow? *) fun mk_insts (env,flexflexes) = (Vartab.dest (Envir.type_env env), Envir.alist_of env); val initenv = Envir.Envir {asol = Vartab.empty, iTs = typinsttab, maxidx = ix2}; val useq = (Unify.unifiers (sgn,initenv,[a])) - handle Library.LIST _ => Seq.empty - | Term.TERM _ => Seq.empty; + handle UnequalLengths => Seq.empty + | Term.TERM _ => Seq.empty; fun clean_unify' useq () = (case (Seq.pull useq) of NONE => NONE | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) - handle Library.LIST _ => NONE + handle UnequalLengths => NONE | Term.TERM _ => NONE; in (Seq.make (clean_unify' useq)) @@ -325,7 +325,7 @@ (* cunning version *) -(* This version avoids name conflicts that might be intorduced by +(* This version avoids name conflicts that might be introduced by unskolemisation, and returns a list of (string * Term.typ) * string, where the outer string is the original name, and the inner string is the new name, and the type is the type of the free variable that was @@ -340,7 +340,7 @@ renamedn :: names)) end | (SOME ((renamedn, _), _)) => (renamedn, L) end - handle LIST _ => (n, L); + handle Fail _ => (n, L); fun unsk (L,f) (t1 $ t2) = let val (L', t1') = unsk (L, I) t1 @@ -365,7 +365,7 @@ let (n,ty) = (Term.dest_Free fv) in SOME (fv, Free (Syntax.dest_skolem n, ty)) handle LIST _ => NONE end - val substfrees = mapfilter fmap (Term.term_frees t) + val substfrees = List.mapPartial fmap (Term.term_frees t) in Term.subst_free substfrees t end; *) @@ -379,7 +379,7 @@ (* get the frees in a term that are of atomic type, ie non-functions *) val atomic_frees_of_term = - filter (is_atomic_typ o snd) + List.filter (is_atomic_typ o snd) o map Term.dest_Free o Term.term_frees; @@ -408,8 +408,8 @@ let val lopt = (term_name_eq ah bh l) in - if is_some lopt then - term_name_eq at bt (the lopt) + if isSome lopt then + term_name_eq at bt (valOf lopt) else NONE end @@ -438,12 +438,12 @@ (* checks to see if the term in a name-equivalent member of the list of terms. *) fun get_name_eq_member a [] = NONE | get_name_eq_member a (h :: t) = - if is_some (term_name_eq a h []) then SOME h + if isSome (term_name_eq a h []) then SOME h else get_name_eq_member a t; fun name_eq_member a [] = false | name_eq_member a (h :: t) = - if is_some (term_name_eq a h []) then true + if isSome (term_name_eq a h []) then true else name_eq_member a t; (* true if term is a variable *) @@ -485,7 +485,7 @@ | term_contains1 T (ah $ at) (bh $ bt) = (term_contains1 T ah (bh $ bt)) @ (term_contains1 T at (bh $ bt)) @ - (flat (map (fn inT => (term_contains1 inT at bt)) + (List.concat (map (fn inT => (term_contains1 inT at bt)) (term_contains1 T ah bh))) | term_contains1 T a (bh $ bt) = [] @@ -662,7 +662,7 @@ | SOME _ => nt) | bounds_to_frees_aux T (nt as (ntab, Bound i)) = let - val (s,ty) = Library.nth_elem (i,T) + val (s,ty) = List.nth (T,i) in (ntab, Free (s,ty)) end; @@ -683,7 +683,7 @@ (loose_bnds_to_frees_aux (bnds,vars) a) $ (loose_bnds_to_frees_aux (bnds,vars) b) | loose_bnds_to_frees_aux (bnds,vars) (t as (Bound i)) = - if (bnds <= i) then Free (Library.nth_elem (i - bnds,vars)) + if (bnds <= i) then Free (List.nth (vars,i - bnds)) else t | loose_bnds_to_frees_aux (bnds,vars) t = t; @@ -705,7 +705,7 @@ fun thms_of_prems_in_goal i tm= let - val goal = (nth_elem (i-1,Thm.prems_of tm)) + val goal = (List.nth (Thm.prems_of tm,i-1)) val vars = rev (strip_all_vars goal) val prems = Logic.strip_assums_hyp (strip_all_body goal) val simp_prem_thms = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/args.ML --- a/src/Pure/Isar/args.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/args.ML Thu Mar 03 12:43:01 2005 +0100 @@ -149,7 +149,7 @@ val int = Scan.optional ($$$ "-" >> K ~1) 1 -- nat >> op *; (*read variable name; leading '?' may be omitted if name contains no dot*) -val var = kind (apsome #1 o (fn s => +val var = kind (Option.map #1 o (fn s => SOME (Term.dest_Var (Syntax.read_var s)) handle _ => SOME (Term.dest_Var (Syntax.read_var ("?" ^ s))) handle _ => NONE)); @@ -214,7 +214,7 @@ ((Scan.repeat1 (Scan.repeat1 (atom_arg blk) || paren_args "(" ")" args || - paren_args "[" "]" args)) >> flat) x; + paren_args "[" "]" args)) >> List.concat) x; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/attrib.ML --- a/src/Pure/Isar/attrib.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/attrib.ML Thu Mar 03 12:43:01 2005 +0100 @@ -135,7 +135,7 @@ val thm_sel = Args.parens (Scan.pass () (Args.enum1 "," (Scan.lift ( Args.nat --| Args.$$$ "-" -- Args.nat >> op upto - || Args.nat >> single)) >> flat)); + || Args.nat >> single)) >> List.concat)); fun gen_thm get attrib app = Scan.depend (fn st => Args.name -- Scan.option thm_sel -- Args.opt_attribs >> @@ -143,11 +143,11 @@ val global_thm = gen_thm PureThy.get_thm global_attribute Thm.apply_attributes; val global_thms = gen_thm PureThy.get_thms global_attribute Thm.applys_attributes; -val global_thmss = Scan.repeat global_thms >> flat; +val global_thmss = Scan.repeat global_thms >> List.concat; val local_thm = gen_thm ProofContext.get_thm local_attribute' Thm.apply_attributes; val local_thms = gen_thm ProofContext.get_thms local_attribute' Thm.applys_attributes; -val local_thmss = Scan.repeat local_thms >> flat; +val local_thmss = Scan.repeat local_thms >> List.concat; @@ -218,7 +218,7 @@ "Type instantiations must occur before term instantiations." else (); - val Tinsts = filter has_type_var insts; + val Tinsts = List.filter has_type_var insts; val tinsts = filter_out has_type_var insts; (* Type instantiations first *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/calculation.ML --- a/src/Pure/Isar/calculation.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/calculation.ML Thu Mar 03 12:43:01 2005 +0100 @@ -178,7 +178,7 @@ | SOME calc => (false, Seq.map single (combine (calc @ facts)))); in err_if state (initial andalso final) "No calculation yet"; - err_if state (initial andalso is_some opt_rules) "Initial calculation -- no rules to be given"; + err_if state (initial andalso isSome opt_rules) "Initial calculation -- no rules to be given"; calculations |> Seq.map (fn calc => (print (Proof.context_of state) calc; state |> maintain_calculation final calc)) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/constdefs.ML --- a/src/Pure/Isar/constdefs.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/constdefs.ML Thu Mar 03 12:43:01 2005 +0100 @@ -40,12 +40,12 @@ val ctxt = ProofContext.init thy |> ProofContext.add_fixes - (flat (map (fn (As, T) => map (fn A => (A, T, NONE)) As) structs)); + (List.concat (map (fn (As, T) => map (fn A => (A, T, NONE)) As) structs)); val (ctxt', d, mx) = (case decl of NONE => (ctxt, NONE, Syntax.NoSyn) | SOME (x, raw_T, mx) => let val x' = Syntax.const_name x mx and mx' = Syntax.fix_mixfix x mx; - val T = apsome (prep_typ ctxt) raw_T; + val T = Option.map (prep_typ ctxt) raw_T; in (ProofContext.add_fixes_liberal [(x', T, SOME mx')] ctxt, SOME x', mx') end); val prop = prep_term ctxt' raw_prop; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/context_rules.ML --- a/src/Pure/Isar/context_rules.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/context_rules.ML Thu Mar 03 12:43:01 2005 +0100 @@ -112,8 +112,8 @@ fun print_rules prt x (Rules {rules, ...}) = let fun prt_kind (i, b) = - Pretty.big_list (the (assoc (kind_names, (i, b))) ^ ":") - (mapfilter (fn (_, (k, th)) => if k = (i, b) then SOME (prt x th) else NONE) + Pretty.big_list (valOf (assoc (kind_names, (i, b))) ^ ":") + (List.mapPartial (fn (_, (k, th)) => if k = (i, b) then SOME (prt x th) else NONE) (sort (Int.compare o pairself fst) rules)); in Pretty.writeln (Pretty.chunks (map prt_kind rule_kinds)) end; @@ -136,7 +136,7 @@ val rules = gen_merge_lists' (fn ((_, (k1, th1)), (_, (k2, th2))) => k1 = k2 andalso Drule.eq_thm_prop (th1, th2)) rules1 rules2; val next = ~ (length rules); - val netpairs = foldl (fn (nps, (n, (w, ((i, b), th)))) => + val netpairs = Library.foldl (fn (nps, (n, (w, ((i, b), th)))) => map_nth_elem i (curry insert_tagged_brl ((w, n), (b, th))) nps) (empty_netpairs, next upto ~1 ~~ rules); in make_rules (next - 1) rules netpairs wrappers end; @@ -203,7 +203,7 @@ fun gen_wrap which ctxt = let val Rules {wrappers, ...} = LocalRules.get ctxt - in fn tac => foldr (fn ((w, _), t) => w t) (which wrappers, tac) end; + in fn tac => Library.foldr (fn ((w, _), t) => w t) (which wrappers, tac) end; val Swrap = gen_wrap #1; val wrap = gen_wrap #2; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/induct_attrib.ML --- a/src/Pure/Isar/induct_attrib.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/induct_attrib.ML Thu Mar 03 12:43:01 2005 +0100 @@ -66,10 +66,10 @@ val vars_of = rev o rev_vars_of; -fun first_var thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle LIST _ => +fun first_var thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle Empty => raise THM ("No variables in first premise of rule", 0, [thm]); -fun last_var thm = mk_var (hd (rev_vars_of (Thm.concl_of thm))) handle LIST _ => +fun last_var thm = mk_var (hd (rev_vars_of (Thm.concl_of thm))) handle Empty => raise THM ("No variables in conclusion of rule", 0, [thm]); end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/isar_output.ML --- a/src/Pure/Isar/isar_output.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/isar_output.ML Thu Mar 03 12:43:01 2005 +0100 @@ -66,7 +66,7 @@ else warning ("Redefined antiquotation output " ^ kind ^ ": " ^ quote name); Symtab.update ((name, x), tab)); -fun add_items kind xs tab = foldl (add_item kind) (tab, xs); +fun add_items kind xs tab = Library.foldl (add_item kind) (tab, xs); in @@ -200,7 +200,7 @@ T.is_kind T.Command tk andalso T.val_of tk mem !hidden_commands | is_hidden _ = false; -fun filter_newlines toks = (case mapfilter +fun filter_newlines toks = (case List.mapPartial (fn (tk, i) => if is_newline tk then SOME tk else NONE) toks of [] => [] | [tk] => [tk] | _ :: toks' => toks'); @@ -209,7 +209,7 @@ ((if !hide_commands andalso exists (is_hidden o fst) toks then [] else if !hide_commands andalso is_proof state then if is_proof state' then [] else filter_newlines toks - else mapfilter (fn (tk, i) => + else List.mapPartial (fn (tk, i) => if i > ! interest_level then NONE else SOME tk) toks) |> map (apsnd (eval_antiquote lex state')) |> Latex.output_tokens diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/isar_syn.ML Thu Mar 03 12:43:01 2005 +0100 @@ -214,7 +214,7 @@ val declareP = OuterSyntax.command "declare" "declare theorems (improper)" K.thy_script - (P.locale_target -- (P.and_list1 P.xthms1 >> flat) + (P.locale_target -- (P.and_list1 P.xthms1 >> List.concat) >> (Toplevel.theory o uncurry IsarThy.declare_theorems)); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/isar_thy.ML --- a/src/Pure/Isar/isar_thy.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/isar_thy.ML Thu Mar 03 12:43:01 2005 +0100 @@ -247,8 +247,8 @@ fun declare_theorems opt_loc args = #1 o smart_theorems "" opt_loc [(("", []), args)]; -fun apply_theorems args = apsnd (flat o map snd) o theorems "" [(("", []), args)]; -fun apply_theorems_i args = apsnd (flat o map snd) o theorems_i "" [(("", []), args)]; +fun apply_theorems args = apsnd (List.concat o map snd) o theorems "" [(("", []), args)]; +fun apply_theorems_i args = apsnd (List.concat o map snd) o theorems_i "" [(("", []), args)]; end; @@ -304,7 +304,7 @@ local fun prt_facts ctxt = - flat o (separate [Pretty.fbrk, Pretty.str "and "]) o + List.concat o (separate [Pretty.fbrk, Pretty.str "and "]) o map (single o ProofContext.pretty_fact ctxt); fun pretty_results ctxt ((kind, ""), facts) = @@ -503,7 +503,7 @@ fun proof''' f = Toplevel.proof' (f o cond_print_calc); fun gen_calc get f args prt state = - f (apsome (fn srcs => get srcs state) args) prt state; + f (Option.map (fn srcs => get srcs state) args) prt state; in diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/locale.ML --- a/src/Pure/Isar/locale.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/locale.ML Thu Mar 03 12:43:01 2005 +0100 @@ -194,7 +194,7 @@ val sign = ProofContext.sign_of ctxt; fun prt_id (name, parms) = [Pretty.block (Pretty.breaks (map Pretty.str (cond_extern sign name :: parms)))]; - val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids)); + val prt_ids = List.concat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids)); val err_msg = if forall (equal "" o #1) ids then msg else msg ^ "\n" ^ Pretty.string_of (Pretty.block @@ -210,7 +210,7 @@ (* renaming *) -fun rename ren x = if_none (assoc_string (ren, x)) x; +fun rename ren x = getOpt (assoc_string (ren, x), x); fun rename_term ren (Free (x, T)) = Free (rename ren x, T) | rename_term ren (t $ u) = rename_term ren t $ rename_term ren u @@ -221,7 +221,7 @@ let val {sign, hyps, prop, maxidx, ...} = Thm.rep_thm th; val cert = Thm.cterm_of sign; - val (xs, Ts) = Library.split_list (foldl Term.add_frees ([], prop :: hyps)); + val (xs, Ts) = Library.split_list (Library.foldl Term.add_frees ([], prop :: hyps)); val xs' = map (rename ren) xs; fun cert_frees names = map (cert o Free) (names ~~ Ts); fun cert_vars names = map (cert o Var o apfst (rpair (maxidx + 1))) (names ~~ Ts); @@ -264,7 +264,7 @@ (* type instantiation *) fun inst_type [] T = T - | inst_type env T = Term.map_type_tfree (fn v => if_none (assoc (env, v)) (TFree v)) T; + | inst_type env T = Term.map_type_tfree (fn v => getOpt (assoc (env, v), TFree v)) T; fun inst_term [] t = t | inst_term env t = Term.map_term_types (inst_type env) t; @@ -276,8 +276,8 @@ val cert = Thm.cterm_of sign; val certT = Thm.ctyp_of sign; val {hyps, prop, maxidx, ...} = Thm.rep_thm th; - val tfrees = foldr Term.add_term_tfree_names (prop :: hyps, []); - val env' = filter (fn ((a, _), _) => a mem_string tfrees) env; + val tfrees = Library.foldr Term.add_term_tfree_names (prop :: hyps, []); + val env' = List.filter (fn ((a, _), _) => a mem_string tfrees) env; in if null env' then th else @@ -285,13 +285,13 @@ |> Drule.implies_intr_list (map cert hyps) |> Drule.tvars_intr_list (map (#1 o #1) env') |> (fn (th', al) => th' |> - Thm.instantiate ((map (fn ((a, _), T) => (the (assoc (al, a)), certT T)) env'), [])) + Thm.instantiate ((map (fn ((a, _), T) => (valOf (assoc (al, a)), certT T)) env'), [])) |> (fn th'' => Drule.implies_elim_list th'' (map (Thm.assume o cert o inst_term env') hyps)) end; fun inst_elem _ env (Fixes fixes) = - Fixes (map (fn (x, T, mx) => (x, apsome (inst_type env) T, mx)) fixes) + Fixes (map (fn (x, T, mx) => (x, Option.map (inst_type env) T, mx)) fixes) | inst_elem _ env (Assumes asms) = Assumes (map (apsnd (map (fn (t, (ps, qs)) => (inst_term env t, (map (inst_term env) ps, map (inst_term env) qs))))) asms) | inst_elem _ env (Defines defs) = Defines (map (apsnd (fn (t, ps) => @@ -310,7 +310,7 @@ fun frozen_tvars ctxt Ts = let - val tvars = rev (foldl Term.add_tvarsT ([], Ts)); + val tvars = rev (Library.foldl Term.add_tvarsT ([], Ts)); val tfrees = map TFree (Term.invent_names (ProofContext.used_types ctxt) "'a" (length tvars) ~~ map #2 tvars); in map #1 tvars ~~ tfrees end; @@ -328,17 +328,17 @@ handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], [])) | unify (env, _) = env; - val (unifier, _) = foldl unify ((Vartab.empty, maxidx''), Ts' ~~ Us'); - val Vs = map (apsome (Envir.norm_type unifier)) Us'; - val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (mapfilter I Vs)); - in map (apsome (Envir.norm_type unifier')) Vs end; + val (unifier, _) = Library.foldl unify ((Vartab.empty, maxidx''), Ts' ~~ Us'); + val Vs = map (Option.map (Envir.norm_type unifier)) Us'; + val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (List.mapPartial I Vs)); + in map (Option.map (Envir.norm_type unifier')) Vs end; -fun params_of elemss = gen_distinct eq_fst (flat (map (snd o fst) elemss)); -fun params_of' elemss = gen_distinct eq_fst (flat (map (snd o fst o fst) elemss)); +fun params_of elemss = gen_distinct eq_fst (List.concat (map (snd o fst) elemss)); +fun params_of' elemss = gen_distinct eq_fst (List.concat (map (snd o fst o fst) elemss)); (* CB: param_types has the following type: ('a * 'b option) list -> ('a * 'b) list *) -fun param_types ps = mapfilter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps; +fun param_types ps = List.mapPartial (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps; (* flatten expressions *) @@ -353,7 +353,7 @@ fun unique_parms ctxt elemss = let val param_decls = - flat (map (fn (((name, (ps, qs)), _), _) => map (rpair (name, ps)) qs) elemss) + List.concat (map (fn (((name, (ps, qs)), _), _) => map (rpair (name, ps)) qs) elemss) |> Symtab.make_multi |> Symtab.dest; in (case find_first (fn (_, ids) => length ids > 1) param_decls of @@ -377,7 +377,7 @@ fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S)); fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps); - val parms = fixed_parms @ flat (map varify_parms idx_parmss); + val parms = fixed_parms @ List.concat (map varify_parms idx_parmss); fun unify T ((env, maxidx), U) = Type.unify tsig (env, maxidx) (U, T) @@ -386,17 +386,17 @@ in raise TYPE ("unify_parms: failed to unify types " ^ prt U ^ " and " ^ prt T, [U, T], []) end - fun unify_list (envir, T :: Us) = foldl (unify T) (envir, Us) + fun unify_list (envir, T :: Us) = Library.foldl (unify T) (envir, Us) | unify_list (envir, []) = envir; - val (unifier, _) = foldl unify_list + val (unifier, _) = Library.foldl unify_list ((Vartab.empty, maxidx), map #2 (Symtab.dest (Symtab.make_multi parms))); val parms' = map (apsnd (Envir.norm_type unifier)) (gen_distinct eq_fst parms); val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms')); fun inst_parms (i, ps) = - foldr Term.add_typ_tfrees (mapfilter snd ps, []) - |> mapfilter (fn (a, S) => + Library.foldr Term.add_typ_tfrees (List.mapPartial snd ps, []) + |> List.mapPartial (fn (a, S) => let val T = Envir.norm_type unifier' (TVar ((a, i), S)) in if T = TFree (a, S) then NONE else SOME ((a, S), T) end) in map inst_parms idx_parmss end; @@ -411,7 +411,7 @@ let val envs = unify_parms ctxt fixed_parms (map (#2 o #1 o #1) elemss); fun inst ((((name, ps), axs), elems), env) = - (((name, map (apsnd (apsome (inst_type env))) ps), axs), + (((name, map (apsnd (Option.map (inst_type env))) ps), axs), map (inst_elem ctxt env) elems); in map inst (elemss ~~ envs) end; @@ -421,7 +421,7 @@ let val envs = unify_parms ctxt fixed_parms (map (#2 o #1 o #1) elemss); fun inst ((((name, ps), axs), elems), env) = - (((name, map (apsnd (apsome (inst_type env))) ps), + (((name, map (apsnd (Option.map (inst_type env))) ps), map (inst_thm ctxt env) axs), map (inst_elem ctxt env) elems); in map inst (elemss ~~ envs) end; @@ -459,8 +459,8 @@ val ids_ax = if top then snd (foldl_map (fn (axs, ((name, parms), _)) => let val {elems, ...} = the_locale thy name; - val ts = flat (mapfilter (fn (Assumes asms, _) => - SOME (flat (map (map #1 o #2) asms)) | _ => NONE) elems); + val ts = List.concat (List.mapPartial (fn (Assumes asms, _) => + SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems); val (axs1, axs2) = splitAt (length ts, axs); in (axs2, ((name, parms), (ps, axs1))) end) (axioms, ids'')) else ids''; @@ -471,10 +471,10 @@ val ren = renaming xs parms' handle ERROR_MESSAGE msg => err_in_locale' ctxt msg ids'; val ids'' = gen_distinct eq_fst (map (rename_parms top ren) ids'); - val parms'' = distinct (flat (map (#2 o #1) ids'')); + val parms'' = distinct (List.concat (map (#2 o #1) ids'')); in (ids'', parms'') end | identify top (Merge es) = - foldl (fn ((ids, parms), e) => let + Library.foldl (fn ((ids, parms), e) => let val (ids', parms') = identify top e in (gen_merge_lists eq_fst ids ids', merge_lists parms parms') end) @@ -508,7 +508,7 @@ elemss; fun inst_ax th = let val {hyps, prop, ...} = Thm.rep_thm th; - val ps = map (apsnd SOME) (foldl Term.add_frees ([], prop :: hyps)); + val ps = map (apsnd SOME) (Library.foldl Term.add_frees ([], prop :: hyps)); val [env] = unify_parms ctxt all_params [ps]; val th' = inst_thm ctxt env th; in th' end; @@ -534,7 +534,7 @@ ((ctxt |> ProofContext.add_fixes fixes, axs), []) | activate_elem _ ((ctxt, axs), Assumes asms) = let - val ts = flat (map (map #1 o #2) asms); + val ts = List.concat (map (map #1 o #2) asms); val (ps,qs) = splitAt (length ts, axs); val (ctxt', _) = ctxt |> ProofContext.fix_frees ts @@ -560,7 +560,7 @@ fun activate_elemss prep_facts = foldl_map (fn (ctxt, (((name, ps), axs), raw_elems)) => let val elems = map (prep_facts ctxt) raw_elems; - val (ctxt', res) = apsnd flat (activate_elems (((name, ps), axs), elems) ctxt); + val (ctxt', res) = apsnd List.concat (activate_elems (((name, ps), axs), elems) ctxt); in (ctxt', (((name, ps), elems), res)) end); in @@ -577,7 +577,7 @@ the internal/external markers from elemss. *) fun activate_facts prep_facts arg = - apsnd (apsnd flat o Library.split_list) (activate_elemss prep_facts arg); + apsnd (apsnd List.concat o Library.split_list) (activate_elemss prep_facts arg); end; @@ -670,7 +670,7 @@ fun declare_int_elem (ctxt, Fixes fixes) = (ctxt |> ProofContext.add_fixes (map (fn (x, T, mx) => - (x, apsome (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes), []) + (x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes), []) | declare_int_elem (ctxt, _) = (ctxt, []); fun declare_ext_elem prep_fixes (ctxt, Fixes fixes) = @@ -700,7 +700,7 @@ ProofContext.declare_terms (map Free fixed_params) ctxt; val int_elemss = raw_elemss - |> mapfilter (fn (id, Int es) => SOME (id, es) | _ => NONE) + |> List.mapPartial (fn (id, Int es) => SOME (id, es) | _ => NONE) |> unify_elemss ctxt_with_fixed fixed_params; val (_, raw_elemss') = foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x) @@ -747,14 +747,14 @@ fun eval_text _ _ _ (text, Fixes _) = text | eval_text _ _ is_ext ((((exts, exts'), (ints, ints')), (xs, env, defs)), Assumes asms) = let - val ts = flat (map (map #1 o #2) asms); + val ts = List.concat (map (map #1 o #2) asms); val ts' = map (norm_term env) ts; val spec' = if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) else ((exts, exts'), (ints @ ts, ints' @ ts')); - in (spec', (foldl Term.add_frees (xs, ts'), env, defs)) end + in (spec', (Library.foldl Term.add_frees (xs, ts'), env, defs)) end | eval_text ctxt (id, _) _ ((spec, binds), Defines defs) = - (spec, foldl (bind_def ctxt id) (binds, map (#1 o #2) defs)) + (spec, Library.foldl (bind_def ctxt id) (binds, map (#1 o #2) defs)) | eval_text _ _ _ (text, Notes _) = text; (* CB: for finish_elems (Ext) *) @@ -798,7 +798,7 @@ fun finish_elems ctxt parms _ (text, ((id, Int e), _)) = let val [(id', es)] = unify_elemss ctxt parms [(id, e)]; - val text' = foldl (eval_text ctxt id' false) (text, es); + val text' = Library.foldl (eval_text ctxt id' false) (text, es); in (text', (id', map Int es)) end | finish_elems ctxt parms do_close (text, ((id, Ext e), [propp])) = let @@ -830,8 +830,8 @@ the fixes elements in raw_elemss, raw_proppss contains assumptions and definitions from the external elements in raw_elemss. *) - val raw_propps = map flat raw_proppss; - val raw_propp = flat raw_propps; + val raw_propps = map List.concat raw_proppss; + val raw_propp = List.concat raw_propps; (* CB: add type information from fixed_params to context (declare_terms) *) (* CB: process patterns (conclusion and external elements only) *) @@ -840,7 +840,7 @@ (* CB: add type information from conclusion and external elements to context *) - val ctxt = ProofContext.declare_terms (flat (map (map fst) all_propp)) ctxt; + val ctxt = ProofContext.declare_terms (List.concat (map (map fst) all_propp)) ctxt; (* CB: resolve schematic variables (patterns) in conclusion and external elements. *) @@ -929,7 +929,7 @@ val (import_ids, raw_import_elemss) = flatten (context, prep_expr sign) ([], Expr import); (* CB: normalise "includes" among elements *) - val raw_elemss = flat (#2 ((foldl_map (flatten (context, prep_expr sign)) + val raw_elemss = List.concat (#2 ((foldl_map (flatten (context, prep_expr sign)) (import_ids, elements)))); (* CB: raw_import_elemss @ raw_elemss is the normalised list of context elements obtained from import and elements. *) @@ -943,8 +943,8 @@ val (ctxt, (elemss, _)) = activate_facts prep_facts (import_ctxt, qs); val stmt = gen_distinct Term.aconv - (flat (map (fn ((_, axs), _) => - flat (map (#hyps o Thm.rep_thm) axs)) qs)); + (List.concat (map (fn ((_, axs), _) => + List.concat (map (#hyps o Thm.rep_thm) axs)) qs)); val cstmt = map (cterm_of sign) stmt; in ((((import_ctxt, import_elemss), (ctxt, elemss)), (parms, spec, defs)), (cstmt, concl)) @@ -956,7 +956,7 @@ fun gen_statement prep_locale prep_ctxt raw_locale elems concl ctxt = let val thy = ProofContext.theory_of ctxt; - val locale = apsome (prep_locale (Theory.sign_of thy)) raw_locale; + val locale = Option.map (prep_locale (Theory.sign_of thy)) raw_locale; val (target_stmt, fixed_params, import) = (case locale of NONE => ([], [], empty) | SOME name => @@ -1036,7 +1036,7 @@ (* process parameters: match their types with those of arguments *) val def_names = map (fn (Free (s, _), _) => s) env; - val (defined, assumed) = partition + val (defined, assumed) = List.partition (fn (s, _) => s mem def_names) fixed_params; val cassumed = map (cert o Free) assumed; val cdefined = map (cert o Free) defined; @@ -1044,21 +1044,21 @@ val param_types = map snd assumed; val v_param_types = map Type.varifyT param_types; val arg_types = map Term.fastype_of args; - val Tenv = foldl (Type.typ_match tsig) + val Tenv = Library.foldl (Type.typ_match tsig) (Vartab.empty, v_param_types ~~ arg_types) - handle Library.LIST "~~" => error "Number of parameters does not \ + handle UnequalLengths => error "Number of parameters does not \ \match number of arguments of chained fact"; (* get their sorts *) - val tfrees = foldr Term.add_typ_tfrees (param_types, []); + val tfrees = Library.foldr Term.add_typ_tfrees (param_types, []); val Tenv' = map - (fn ((a, i), T) => ((a, the (assoc_string (tfrees, a))), T)) + (fn ((a, i), T) => ((a, valOf (assoc_string (tfrees, a))), T)) (Vartab.dest Tenv); (* process (internal) elements *) fun inst_type [] T = T | inst_type env T = - Term.map_type_tfree (fn v => if_none (assoc (env, v)) (TFree v)) T; + Term.map_type_tfree (fn v => getOpt (assoc (env, v), TFree v)) T; fun inst_term [] t = t | inst_term env t = Term.map_term_types (inst_type env) t; @@ -1076,8 +1076,8 @@ val cert = Thm.cterm_of sign; val certT = Thm.ctyp_of sign; val {hyps, prop, maxidx, ...} = Thm.rep_thm th; - val tfrees = foldr Term.add_term_tfree_names (prop :: hyps, []); - val Tenv' = filter (fn ((a, _), _) => a mem_string tfrees) Tenv; + val tfrees = Library.foldr Term.add_term_tfree_names (prop :: hyps, []); + val Tenv' = List.filter (fn ((a, _), _) => a mem_string tfrees) Tenv; in if null Tenv' then th else @@ -1086,7 +1086,7 @@ |> Drule.tvars_intr_list (map (#1 o #1) Tenv') |> (fn (th', al) => th' |> Thm.instantiate ((map (fn ((a, _), T) => - (the (assoc (al, a)), certT T)) Tenv'), [])) + (valOf (assoc (al, a)), certT T)) Tenv'), [])) |> (fn th'' => Drule.implies_elim_list th'' (map (Thm.assume o cert o inst_term Tenv') hyps)) end; @@ -1096,7 +1096,7 @@ (* not all axs are normally applicable *) val hyps = #hyps (rep_thm thm); val ass = map (fn ax => (prop_of ax, ax)) axioms; - val axs' = foldl (fn (axs, hyp) => + val axs' = Library.foldl (fn (axs, hyp) => (case gen_assoc (op aconv) (ass, hyp) of NONE => axs | SOME ax => axs @ [ax])) ([], hyps); val thm' = Drule.satisfy_hyps axs' thm; @@ -1142,9 +1142,9 @@ end | inst_elem (ctxt, (Int _)) = ctxt; - fun inst_elems (ctxt, (id, elems)) = foldl inst_elem (ctxt, elems); + fun inst_elems (ctxt, (id, elems)) = Library.foldl inst_elem (ctxt, elems); - fun inst_elemss ctxt elemss = foldl inst_elems (ctxt, elemss); + fun inst_elemss ctxt elemss = Library.foldl inst_elems (ctxt, elemss); (* main part *) @@ -1161,7 +1161,7 @@ let val thy_ctxt = ProofContext.init thy; val (((_, import_elemss), (ctxt, elemss)), _) = read_context import body thy_ctxt; - val all_elems = flat (map #2 (import_elemss @ elemss)); + val all_elems = List.concat (map #2 (import_elemss @ elemss)); val prt_typ = Pretty.quote o ProofContext.pretty_typ ctxt; val prt_term = Pretty.quote o ProofContext.pretty_term ctxt; @@ -1179,7 +1179,7 @@ fun prt_asm ((a, _), ts) = Pretty.block (prt_name a @ Pretty.breaks (map (prt_term o fst) ts)); fun prt_def ((a, _), (t, _)) = Pretty.block (prt_name a @ [prt_term t]); fun prt_fact ((a, _), ths) = Pretty.block - (prt_name a @ Pretty.breaks (map prt_thm (flat (map fst ths)))); + (prt_name a @ Pretty.breaks (map prt_thm (List.concat (map fst ths)))); fun items _ [] = [] | items prfx (x :: xs) = Pretty.block [Pretty.str prfx, Pretty.brk 1, x] :: items " and" xs; @@ -1306,9 +1306,9 @@ val (body, bodyT, body_eq) = atomize_spec sign norm_ts; val env = Term.add_term_free_names (body, []); - val xs = filter (fn (x, _) => x mem_string env) parms; + val xs = List.filter (fn (x, _) => x mem_string env) parms; val Ts = map #2 xs; - val extraTs = (Term.term_tfrees body \\ foldr Term.add_typ_tfrees (Ts, [])) + val extraTs = (Term.term_tfrees body \\ Library.foldr Term.add_typ_tfrees (Ts, [])) |> Library.sort_wrt #1 |> map TFree; val predT = map Term.itselfT extraTs ---> Ts ---> bodyT; @@ -1408,7 +1408,7 @@ let val sign = Theory.sign_of thy; val name = Sign.full_name sign bname; - val _ = conditional (is_some (get_locale thy name)) (fn () => + val _ = conditional (isSome (get_locale thy name)) (fn () => error ("Duplicate definition of locale " ^ quote name)); val thy_ctxt = ProofContext.init thy; @@ -1423,8 +1423,8 @@ fun axiomify axioms elemss = (axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let - val ts = flat (mapfilter (fn (Assumes asms) => - SOME (flat (map (map #1 o #2) asms)) | _ => NONE) elems); + val ts = List.concat (List.mapPartial (fn (Assumes asms) => + SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems); val (axs1, axs2) = splitAt (length ts, axs); in (axs2, ((id, axs1), elems)) end) |> snd; @@ -1437,7 +1437,7 @@ |> note_thmss_qualified "" name facts' |> #1 |> declare_locale name |> put_locale name {predicate = predicate, import = prep_expr sign raw_import, - elems = map (fn e => (e, stamp ())) (flat (map #2 (filter (equal "" o #1 o #1) elemss'))), + elems = map (fn e => (e, stamp ())) (List.concat (map #2 (List.filter (equal "" o #1 o #1) elemss'))), params = (params_of elemss', map #1 (params_of body_elemss))} end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/method.ML --- a/src/Pure/Isar/method.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/method.ML Thu Mar 03 12:43:01 2005 +0100 @@ -152,7 +152,7 @@ (* shuffle *) fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); -fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); +fun defer opt_i = METHOD (K (Tactic.defer_tac (getOpt (opt_i,1)))); (* insert *) @@ -261,7 +261,7 @@ in fun rules_tac ctxt opt_lim = - SELECT_GOAL (DEEPEN (2, if_none opt_lim 20) (intpr_tac ctxt [] 0) 4 1); + SELECT_GOAL (DEEPEN (2, getOpt (opt_lim,20)) (intpr_tac ctxt [] 0) 4 1); end; @@ -281,7 +281,7 @@ let val rules = if not (null arg_rules) then arg_rules - else flat (ContextRules.find_rules false facts goal ctxt) + else List.concat (ContextRules.find_rules false facts goal ctxt) in trace ctxt rules; tac rules facts i end); fun meth tac x = METHOD (HEADGOAL o tac x); @@ -309,7 +309,7 @@ (* assumption *) fun asm_tac ths = - foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac); + Library.foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac); fun assm_tac ctxt = assume_tac APPEND' @@ -334,7 +334,7 @@ (* Separate type and term insts *) fun has_type_var ((x, _), _) = (case Symbol.explode x of "'"::cs => true | cs => false); - val Tinsts = filter has_type_var insts; + val Tinsts = List.filter has_type_var insts; val tinsts = filter_out has_type_var insts; (* Tactic *) fun tac i st = @@ -373,7 +373,7 @@ NONE => types (a, ~1) | some => some) | types' xi = types xi; - fun internal x = is_some (types' (x, ~1)); + fun internal x = isSome (types' (x, ~1)); val used = Term.add_term_tvarnames (Thm.prop_of st $ Thm.prop_of thm, []); val (ts, envT) = ProofContext.read_termTs_schematic ctxt internal types' sorts used (ss ~~ Ts); @@ -596,7 +596,7 @@ fun sect ss = Scan.first (map Scan.lift ss); fun thms ss = Scan.unless (sect ss) Attrib.local_thms; -fun thmss ss = Scan.repeat (thms ss) >> flat; +fun thmss ss = Scan.repeat (thms ss) >> List.concat; fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); @@ -732,7 +732,7 @@ fun proof opt_text state = state |> Proof.assert_backward - |> refine (if_none opt_text default_text) + |> refine (getOpt (opt_text,default_text)) |> Seq.map (Proof.goal_facts (K [])) |> Seq.map Proof.enter_forward; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/net_rules.ML --- a/src/Pure/Isar/net_rules.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/net_rules.ML Thu Mar 03 12:43:01 2005 +0100 @@ -51,7 +51,7 @@ fun merge (Rules {eq, index, rules = rules1, ...}, Rules {rules = rules2, ...}) = let val rules = Library.gen_merge_lists' eq rules1 rules2 - in foldr (uncurry add) (rules, init eq index) end; + in Library.foldr (uncurry add) (rules, init eq index) end; fun delete rule (rs as Rules {eq, index, rules, next, net}) = if not (Library.gen_mem eq (rule, rules)) then rs diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/object_logic.ML --- a/src/Pure/Isar/object_logic.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/object_logic.ML Thu Mar 03 12:43:01 2005 +0100 @@ -48,7 +48,7 @@ fun merge_judgment (SOME x, SOME y) = if x = y then SOME x else error "Attempt to merge different object-logics" - | merge_judgment (j1, j2) = if is_some j1 then j1 else j2; + | merge_judgment (j1, j2) = if isSome j1 then j1 else j2; fun merge ((judgment1, (atomize1, rulify1)), (judgment2, (atomize2, rulify2))) = (merge_judgment (judgment1, judgment2), @@ -116,7 +116,7 @@ val c = judgment_name sg; val aT = TFree ("'a", []); val T = - if_none (Sign.const_type sg c) (aT --> propT) + getOpt (Sign.const_type sg c, aT --> propT) |> Term.map_type_tvar (fn ((x, _), S) => TFree (x, S)); val U = Term.domain_type T handle Match => aT; in Const (c, T) $ Free (x, U) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/obtain.ML Thu Mar 03 12:43:01 2005 +0100 @@ -72,12 +72,12 @@ (*obtain vars*) val (vars_ctxt, vars) = foldl_map prep_vars (Proof.context_of state, raw_vars); - val xs = flat (map fst vars); + val xs = List.concat (map fst vars); val fix_ctxt = vars_ctxt |> ProofContext.fix_i vars; (*obtain asms*) val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms); - val asm_props = flat (map (map fst) proppss); + val asm_props = List.concat (map (map fst) proppss); val asms = map fst raw_asms ~~ proppss; val _ = ProofContext.warn_extra_tfrees fix_ctxt asms_ctxt; @@ -95,9 +95,9 @@ fun occs_var x = Library.get_first (fn t => ProofContext.find_free t (ProofContext.get_skolem fix_ctxt x)) asm_props; val raw_parms = map occs_var xs; - val parms = mapfilter I raw_parms; + val parms = List.mapPartial I raw_parms; val parm_names = - mapfilter (fn (SOME (Free a), x) => SOME (a, x) | _ => NONE) (raw_parms ~~ xs); + List.mapPartial (fn (SOME (Free a), x) => SOME (a, x) | _ => NONE) (raw_parms ~~ xs); val that_prop = Term.list_all_free (map #1 parm_names, Logic.list_implies (asm_props, bound_thesis)) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/outer_parse.ML --- a/src/Pure/Isar/outer_parse.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/outer_parse.ML Thu Mar 03 12:43:01 2005 +0100 @@ -278,7 +278,7 @@ ((Scan.repeat1 (Scan.repeat1 (atom_arg is_symid blk) || paren_args "(" ")" (args is_symid) || - paren_args "[" "]" (args is_symid))) >> flat) x; + paren_args "[" "]" (args is_symid))) >> List.concat) x; val arguments = args T.is_sid false; @@ -298,7 +298,7 @@ val thm_sel = $$$ "(" |-- list1 ( nat --| minus -- nat >> op upto - || nat >> single) --| $$$ ")" >> flat; + || nat >> single) --| $$$ ")" >> List.concat; val xthm = xname -- Scan.option thm_sel -- opt_attribs; val xthms1 = Scan.repeat1 xthm; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/outer_syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -183,7 +183,7 @@ fun get_lexicons () = ! global_lexicons; fun get_parsers () = ! global_parsers; -fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers); +fun get_parser () = Option.map (#2 o #1) o curry Symtab.lookup (! global_parsers); fun is_markup kind name = (case assoc (! global_markups, name) of SOME k => k = kind | NONE => false); @@ -201,7 +201,7 @@ Symtab.update ((name, (((comment, kind), (int_only, parse)), markup)), tab)); fun add_parsers parsers = - (change_parsers (fn tab => foldl add_parser (tab, parsers)); + (change_parsers (fn tab => Library.foldl add_parser (tab, parsers)); change_lexicons (apsnd (fn lex => Scan.extend_lexicon lex (map (fn Parser (name, _, _, _) => Symbol.explode name) parsers)))); @@ -221,7 +221,7 @@ let fun pretty_cmd (name, comment, _, _) = Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; - val (int_cmds, cmds) = partition #4 (dest_parsers ()); + val (int_cmds, cmds) = List.partition #4 (dest_parsers ()); in [Pretty.strs ("syntax keywords:" :: map quote (dest_keywords ())), Pretty.big_list "proper commands:" (map pretty_cmd cmds), @@ -277,7 +277,7 @@ Source.of_list toks |> toplevel_source false true true get_parser |> Source.exhaust - |> map (fn tr => (Toplevel.name_of tr, the (Toplevel.source_of tr), tr)); + |> map (fn tr => (Toplevel.name_of tr, valOf (Toplevel.source_of tr), tr)); (** read theory **) @@ -296,7 +296,7 @@ val pos = Path.position path; val (name', parents, files) = ThyHeader.scan (src, pos); val ml_path = ThyLoad.ml_path name; - val ml_file = if ml andalso is_some (ThyLoad.check_file NONE ml_path) then [ml_path] else []; + val ml_file = if ml andalso isSome (ThyLoad.check_file NONE ml_path) then [ml_path] else []; in if name <> name' then error ("Filename " ^ quote (Path.pack path) ^ diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/proof.ML --- a/src/Pure/Isar/proof.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/proof.ML Thu Mar 03 12:43:01 2005 +0100 @@ -269,7 +269,7 @@ val reset_facts = put_facts NONE; fun these_factss (state, named_factss) = - state |> put_facts (SOME (flat (map snd named_factss))); + state |> put_facts (SOME (List.concat (map snd named_factss))); (* goal *) @@ -414,7 +414,7 @@ fun prt_names names = (case filter_out (equal "") names of [] => "" - | nms => " " ^ enclose "(" ")" (space_implode " and " (take (7, nms) @ + | nms => " " ^ enclose "(" ")" (space_implode " and " (Library.take (7, nms) @ (if length nms > 7 then ["..."] else [])))); fun prt_goal (_, (i, (((kind, names, _), (goal_facts, thm)), _))) = @@ -546,7 +546,7 @@ val {hyps, sign, ...} = Thm.rep_thm raw_th; val tsig = Sign.tsig_of sign; - val casms = flat (map #1 (ProofContext.assumptions_of (context_of state))); + val casms = List.concat (map #1 (ProofContext.assumptions_of (context_of state))); val bad_hyps = Library.gen_rems Term.aconv (hyps, map Thm.term_of casms); val th = raw_th RS Drule.rev_triv_goal; @@ -626,7 +626,7 @@ |> map_context_result (f (map (pair ("", [])) args)) |> (fn (st, named_facts) => let val (_, (_, ((result, (facts, thm)), x))) = find_goal st; - in st |> map_goal I (K ((result, (facts @ flat (map snd named_facts), thm)), x)) end); + in st |> map_goal I (K ((result, (facts @ List.concat (map snd named_facts), thm)), x)) end); in @@ -754,12 +754,12 @@ |> map_context_result (fn ctxt => prepp (ctxt, raw_propp)); val sign' = sign_of state'; - val props = flat propss; + val props = List.concat propss; val cprop = Thm.cterm_of sign' (Logic.mk_conjunction_list props); val goal = Drule.mk_triv_goal cprop; - val tvars = foldr Term.add_term_tvars (props, []); - val vars = foldr Term.add_term_vars (props, []); + val tvars = Library.foldr Term.add_term_tvars (props, []); + val vars = Library.foldr Term.add_term_vars (props, []); in if null tvars then () else raise STATE ("Goal statement contains illegal schematic type variable(s): " ^ @@ -786,7 +786,7 @@ let val init = init_state thy; val (opt_name, view, locale_ctxt, elems_ctxt, propp) = - prep (apsome fst raw_locale) elems (map snd concl) (context_of init); + prep (Option.map fst raw_locale) elems (map snd concl) (context_of init); in init |> open_block @@ -798,7 +798,7 @@ |> setup_goal I ProofContext.bind_propp_schematic_i ProofContext.fix_frees (Theorem {kind = kind, theory_spec = (a, map (snd o fst) concl), - locale_spec = (case raw_locale of NONE => NONE | SOME (_, x) => SOME (the opt_name, x)), + locale_spec = (case raw_locale of NONE => NONE | SOME (_, x) => SOME (valOf opt_name, x)), view = view}) Seq.single true (map (fst o fst) concl) propp end; @@ -863,7 +863,7 @@ val outer_state = state |> close_block; val outer_ctxt = context_of outer_state; - val ts = flat tss; + val ts = List.concat tss; val results = prep_result state ts raw_thm; val resultq = results |> map (ProofContext.export false goal_ctxt outer_ctxt) @@ -903,7 +903,7 @@ val {kind = k, theory_spec = ((name, atts), attss), locale_spec, view = (target_view, elems_view)} = (case kind of Theorem x => x | _ => err_malformed "finish_global" state); - val ts = flat tss; + val ts = List.concat tss; val locale_results = map (ProofContext.export_standard elems_view goal_ctxt locale_ctxt) (prep_result state ts raw_thm); @@ -920,12 +920,12 @@ |> (fn ((thy', ctxt'), res) => if name = "" andalso null loc_atts then thy' else (thy', ctxt') - |> (#1 o #1 o Locale.add_thmss loc [((name, flat (map #2 res)), loc_atts)]))) + |> (#1 o #1 o Locale.add_thmss loc [((name, List.concat (map #2 res)), loc_atts)]))) |> Locale.smart_note_thmss k locale_spec ((names ~~ attss) ~~ map (single o Thm.no_attributes) (Library.unflat tss results)) |> (fn (thy, res) => (thy, res) |>> (#1 o Locale.smart_note_thmss k locale_spec - [((name, atts), [(flat (map #2 res), [])])])); + [((name, atts), [(List.concat (map #2 res), [])])])); in (theory', ((k, name), results')) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/proof_context.ML Thu Mar 03 12:43:01 2005 +0100 @@ -221,11 +221,11 @@ val fixed_names_of = map #2 o fixes_of; fun is_fixed ctxt x = exists (equal x o #2) (fixes_of ctxt); fun is_known (ctxt as Context {defs = (types, _, _, _), ...}) x = - is_some (Vartab.lookup (types, (x, ~1))) orelse is_fixed ctxt x; + isSome (Vartab.lookup (types, (x, ~1))) orelse is_fixed ctxt x; fun type_occs (Context {defs = (_, _, _, tab), ...}) = tab; fun assumptions_of (Context {asms = ((asms, _), _), ...}) = asms; -fun prems_of (Context {asms = ((_, prems), _), ...}) = flat (map #2 prems); +fun prems_of (Context {asms = ((_, prems), _), ...}) = List.concat (map #2 prems); @@ -398,9 +398,9 @@ map_context (fn (thy, (syn, structs, mixfixed), data, asms, binds, thms, cases, defs, delta, delta_count) => let val is_logtype = Sign.is_logtype (Theory.sign_of thy); - val structs' = structs @ mapfilter prep_struct decls; - val mxs = mapfilter prep_mixfix decls; - val (fixed, mxs_output) = Library.split_list (mapfilter prep_mixfix' decls); + val structs' = structs @ List.mapPartial prep_struct decls; + val mxs = List.mapPartial prep_mixfix decls; + val (fixed, mxs_output) = Library.split_list (List.mapPartial prep_mixfix' decls); val trs = map fixed_tr fixed; val syn' = syn |> Syntax.extend_const_gram is_logtype ("", false) mxs_output @@ -489,7 +489,7 @@ (* internalize Skolem constants *) fun lookup_skolem ctxt x = assoc (fixes_of ctxt, x); -fun get_skolem ctxt x = if_none (lookup_skolem ctxt x) x; +fun get_skolem ctxt x = getOpt (lookup_skolem ctxt x, x); fun no_skolem internal ctxt x = if can Syntax.dest_skolem x then @@ -697,7 +697,7 @@ val ins_occs = foldl_term_types (fn t => foldl_atyps (fn (tab, TFree (x, _)) => Symtab.update_multi ((x, t), tab) | (tab, _) => tab)); -fun ins_skolem def_ty = foldr +fun ins_skolem def_ty = Library.foldr (fn ((x, x'), types) => (case def_ty x' of SOME T => Vartab.update (((x, ~1), T), types) @@ -721,8 +721,8 @@ in fun declare_term t ctxt = declare_occ (ctxt, t); -fun declare_terms ts ctxt = foldl declare_occ (ctxt, ts); -fun declare_terms_syntax ts ctxt = foldl declare_syn (ctxt, ts); +fun declare_terms ts ctxt = Library.foldl declare_occ (ctxt, ts); +fun declare_terms_syntax ts ctxt = Library.foldl declare_syn (ctxt, ts); end; @@ -742,8 +742,8 @@ val extras = Library.gen_rems Library.eq_fst (Symtab.dest occ2, Symtab.dest occ1) - |> map (fn (a, ts) => map (pair a) (mapfilter (try (#1 o Term.dest_Free)) ts)) |> flat - |> mapfilter (fn (a, x) => + |> map (fn (a, ts) => map (pair a) (List.mapPartial (try (#1 o Term.dest_Free)) ts)) |> List.concat + |> List.mapPartial (fn (a, x) => (case def_type ctxt1 false (x, ~1) of NONE => SOME (a, x) | SOME T => if known_tfree a T then NONE else SOME (a, x))); val tfrees = map #1 extras |> Library.sort_strings |> Library.unique_strings; @@ -766,14 +766,14 @@ val occs_inner = type_occs inner; val occs_outer = type_occs outer; fun add (gen, a) = - if is_some (Symtab.lookup (occs_outer, a)) orelse + if isSome (Symtab.lookup (occs_outer, a)) orelse exists still_fixed (Symtab.lookup_multi (occs_inner, a)) then gen else a :: gen; - in fn tfrees => foldl add ([], tfrees) end; + in fn tfrees => Library.foldl add ([], tfrees) end; fun generalize inner outer ts = let - val tfrees = generalize_tfrees inner outer (foldr Term.add_term_tfree_names (ts, [])); + val tfrees = generalize_tfrees inner outer (Library.foldr Term.add_term_tfree_names (ts, [])); fun gen (x, S) = if x mem_string tfrees then TVar ((x, 0), S) else TFree (x, S); in map (Term.map_term_types (Term.map_type_tfree gen)) ts end; @@ -799,7 +799,7 @@ |> Seq.map (fn rule => let val {sign, prop, ...} = Thm.rep_thm rule; - val frees = map (Thm.cterm_of sign) (mapfilter (find_free prop) fixes); + val frees = map (Thm.cterm_of sign) (List.mapPartial (find_free prop) fixes); val tfrees = gen (Term.add_term_tfree_names (prop, [])); in rule @@ -846,8 +846,8 @@ fun del_upd_bind (ctxt, (xi, NONE)) = del_bind (ctxt, xi) | del_upd_bind (ctxt, (xi, SOME t)) = upd_bind (ctxt, (xi, t)); -fun update_binds bs ctxt = foldl upd_bind (ctxt, bs); -fun delete_update_binds bs ctxt = foldl del_upd_bind (ctxt, bs); +fun update_binds bs ctxt = Library.foldl upd_bind (ctxt, bs); +fun delete_update_binds bs ctxt = Library.foldl del_upd_bind (ctxt, bs); (* simult_matches *) @@ -857,7 +857,7 @@ let fun fail () = raise CONTEXT ("Pattern match failed!", ctxt); - val maxidx = foldl (fn (i, (t1, t2)) => + val maxidx = Library.foldl (fn (i, (t1, t2)) => Int.max (i, Int.max (Term.maxidx_of_term t1, Term.maxidx_of_term t2))) (~1, pairs); val envs = Unify.smash_unifiers (sign_of ctxt, Envir.empty maxidx, map swap pairs); (*prefer assignment of variables from patterns*) @@ -871,7 +871,7 @@ if null (map #1 (Envir.alist_of env) inter (map #1 (Drule.vars_of_terms (map #2 pairs)))) then () else fail (); fun norm_bind (xi, t) = if xi mem domain then SOME (xi, Envir.norm_term env t) else NONE; - in mapfilter norm_bind (Envir.alist_of env) end; + in List.mapPartial norm_bind (Envir.alist_of env) end; (* add_binds(_i) *) @@ -879,9 +879,9 @@ local fun gen_bind prep (ctxt, (xi as (x, _), raw_t)) = - ctxt |> delete_update_binds [(xi, apsome (prep ctxt) raw_t)]; + ctxt |> delete_update_binds [(xi, Option.map (prep ctxt) raw_t)]; -fun gen_binds prep binds ctxt = foldl (gen_bind prep) (ctxt, binds); +fun gen_binds prep binds ctxt = Library.foldl (gen_bind prep) (ctxt, binds); in @@ -913,7 +913,7 @@ let val ts = prep_terms ctxt (map snd raw_binds); val (ctxt', binds) = - apsnd flat (foldl_map (prep_bind prep_pats) (ctxt, map fst raw_binds ~~ ts)); + apsnd List.concat (foldl_map (prep_bind prep_pats) (ctxt, map fst raw_binds ~~ ts)); val binds' = if gen then map #1 binds ~~ generalize ctxt' ctxt (map #2 binds) else binds; @@ -945,7 +945,7 @@ in ((ctxt', props), (prop, splitAt(length raw_pats1, pats))) end | prep _ = sys_error "prep_propp"; val ((context', _), propp) = foldl_map (foldl_map prep) - ((context, prep_props schematic context (flat (map (map fst) args))), args); + ((context, prep_props schematic context (List.concat (map (map fst) args))), args); in (context', propp) end; fun matches ctxt (prop, (pats1, pats2)) = @@ -954,7 +954,7 @@ fun gen_bind_propp prepp (ctxt, raw_args) = let val (ctxt', args) = prepp (ctxt, raw_args); - val binds = flat (flat (map (map (matches ctxt')) args)); + val binds = List.concat (List.concat (map (map (matches ctxt')) args)); val propss = map (map #1) args; (*generalize result: context evaluated now, binds added later*) @@ -1051,8 +1051,8 @@ fun app ((ct, ths), (th, attrs)) = let val (ct', th') = Thm.applys_attributes ((ct, get ctxt th), attrs @ more_attrs) in (ct', th' :: ths) end; - val (ctxt', rev_thms) = foldl app ((ctxt, []), ths_attrs); - val thms = flat (rev rev_thms); + val (ctxt', rev_thms) = Library.foldl app ((ctxt, []), ths_attrs); + val thms = List.concat (rev rev_thms); in (ctxt' |> put_thms (name, thms), (name, thms)) end; fun gen_note_thmss get args ctxt = foldl_map (gen_note_thss get) (ctxt, args); @@ -1090,7 +1090,7 @@ fun is_free (Free (x, _)) = not (is_fixed ctxt x) | is_free _ = false; - val extra_frees = filter is_free (term_frees rhs) \\ xs; + val extra_frees = List.filter is_free (term_frees rhs) \\ xs; in conditional (not (forall (is_Bound orf is_free) xs andalso null (duplicates xs))) (fn () => err "Arguments of lhs must be distinct free/bound variables"); @@ -1098,7 +1098,7 @@ err "Element to be defined occurs on rhs"); conditional (not (null extra_frees)) (fn () => err ("Extra free variables on rhs: " ^ commas_quote (map (#1 o dest_Free) extra_frees))); - (c, Term.list_all_free (mapfilter (try Term.dest_Free) xs, eq)) + (c, Term.list_all_free (List.mapPartial (try Term.dest_Free) xs, eq)) end; fun head_of_def cprop = @@ -1134,7 +1134,7 @@ val (ctxt1, propss) = prepp (ctxt, map snd args); val (ctxt2, results) = foldl_map add_assm (ctxt1, map fst args ~~ propss); - val cprops = flat (map #1 results); + val cprops = List.concat (map #1 results); val asmss = map #2 results; val thmss = map #3 results; val ctxt3 = ctxt2 |> map_context @@ -1163,11 +1163,11 @@ else Type.no_tvars T handle TYPE (msg, _, _) => raise CONTEXT (msg, ctxt); val _ = if liberal then () else - (case filter (not o Syntax.is_identifier) (map (no_skolem internal ctxt) xs) of + (case List.filter (not o Syntax.is_identifier) (map (no_skolem internal ctxt) xs) of [] => () | bads => raise CONTEXT ("Bad variable name(s): " ^ commas_quote bads, ctxt)); - val opt_T = apsome (cond_tvars o prep_typ ctxt) raw_T; - val T = if_none opt_T TypeInfer.logicT; + val opt_T = Option.map (cond_tvars o prep_typ ctxt) raw_T; + val T = getOpt (opt_T, TypeInfer.logicT); val ctxt' = ctxt |> declare_terms_syntax (map (fn x => Free (x, T)) xs); in (ctxt', (xs, opt_T)) end; @@ -1192,7 +1192,7 @@ fun err_dups ctxt xs = raise CONTEXT ("Duplicate variable(s): " ^ commas_quote xs, ctxt); val declare = - declare_terms_syntax o mapfilter (fn (_, NONE) => NONE | (x, SOME T) => SOME (Free (x, T))); + declare_terms_syntax o List.mapPartial (fn (_, NONE) => NONE | (x, SOME T) => SOME (Free (x, T))); fun add_vars xs Ts ctxt = let val xs' = Term.variantlist (map Syntax.skolem xs, map #2 (fixes_of ctxt)) in @@ -1213,7 +1213,7 @@ fun gen_fix prep add raw_vars ctxt = let val (ctxt', varss) = foldl_map prep (ctxt, raw_vars); - val vars = rev (flat (map (fn (xs, T) => map (rpair T) xs) varss)); + val vars = rev (List.concat (map (fn (xs, T) => map (rpair T) xs) varss)); val xs = map #1 vars; val Ts = map #2 vars; in @@ -1239,9 +1239,9 @@ fun fix_frees ts ctxt = let - val frees = foldl Term.add_frees ([], ts); + val frees = Library.foldl Term.add_frees ([], ts); fun new (x, T) = if is_fixed ctxt x then NONE else SOME ([x], SOME T); - in fix_direct false (rev (mapfilter new frees)) ctxt end; + in fix_direct false (rev (List.mapPartial new frees)) ctxt end; (*Note: improper use may result in variable capture / dynamic scoping!*) @@ -1263,12 +1263,12 @@ fun prep_case ctxt name xs {fixes, assumes, binds} = let - fun replace (opt_x :: xs) ((y, T) :: ys) = (if_none opt_x y, T) :: replace xs ys + fun replace (opt_x :: xs) ((y, T) :: ys) = (getOpt (opt_x,y), T) :: replace xs ys | replace [] ys = ys | replace (_ :: _) [] = raise CONTEXT ("Too many parameters for case " ^ quote name, ctxt); in - if null (foldr Term.add_typ_tvars (map snd fixes, [])) andalso - null (foldr Term.add_term_vars (flat (map snd assumes), [])) then + if null (Library.foldr Term.add_typ_tvars (map snd fixes, [])) andalso + null (Library.foldr Term.add_term_vars (List.concat (map snd assumes), [])) then {fixes = replace xs fixes, assumes = assumes, binds = map drop_schematic binds} else raise CONTEXT ("Illegal schematic variable(s) in case " ^ quote name, ctxt) end; @@ -1314,7 +1314,7 @@ fun pretty_binds (ctxt as Context {binds, ...}) = let fun prt_bind (xi, (t, T)) = pretty_term ctxt (Logic.mk_equals (Var (xi, T), t)); - val bs = mapfilter smash_option (Vartab.dest binds); + val bs = List.mapPartial smash_option (Vartab.dest binds); in if null bs andalso not (! verbose) then [] else [Pretty.big_list "term bindings:" (map prt_bind bs)] @@ -1327,7 +1327,7 @@ fun pretty_lthms (ctxt as Context {thms = (_, space, tab, _), ...}) = pretty_items (pretty_thm ctxt) "facts:" - (mapfilter smash_option (NameSpace.cond_extern_table space tab)); + (List.mapPartial smash_option (NameSpace.cond_extern_table space tab)); val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms; @@ -1338,8 +1338,8 @@ let fun bind (c, (x, T)) = (c |> fix_i [([x], SOME T)], bind_skolem c [x] (Free (x, T))); val (ctxt', xs) = foldl_map bind (ctxt, fixes); - fun app t = foldl Term.betapply (t, xs); - in (ctxt', (map (apsnd (apsome app)) binds, map (apsnd (map app)) assumes)) end; + fun app t = Library.foldl Term.betapply (t, xs); + in (ctxt', (map (apsnd (Option.map app)) binds, map (apsnd (map app)) assumes)) end; fun pretty_cases (ctxt as Context {cases, ...}) = let @@ -1354,13 +1354,13 @@ fun prt_sect _ _ _ [] = [] | prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s :: - flat (Library.separate sep (map (Library.single o prt) xs))))]; + List.concat (Library.separate sep (map (Library.single o prt) xs))))]; fun prt_case (name, (fixes, (lets, asms))) = Pretty.block (Pretty.fbreaks (Pretty.str (name ^ ":") :: prt_sect "fix" [] (Pretty.str o fst) fixes @ prt_sect "let" [Pretty.str "and"] prt_let - (mapfilter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @ + (List.mapPartial (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @ (if forall (null o #2) asms then [] else prt_sect "assume" [Pretty.str "and"] prt_asm asms))); @@ -1453,7 +1453,7 @@ (case try (transform_error (fn () => get_thms ctxt (name, NONE))) () of NONE => false | SOME ths' => Library.equal_lists Thm.eq_thm (ths, ths')); - in gen_distinct eq_fst (filter valid (FactIndex.find idx index)) end; + in gen_distinct eq_fst (List.filter valid (FactIndex.find idx index)) end; val thms_containing_limit = ref 40; @@ -1469,7 +1469,7 @@ | prt_frees xs = [Pretty.block (Pretty.breaks (Pretty.str "variables" :: map (Pretty.quote o prt_term o Syntax.free) xs))]; - val (cs, xs) = foldl (FactIndex.index_term (is_known ctxt)) + val (cs, xs) = Library.foldl (FactIndex.index_term (is_known ctxt)) (([], []), map (read_term_dummies ctxt) ss); val empty_idx = null cs andalso null xs; @@ -1482,7 +1482,7 @@ PureThy.thms_containing (theory_of ctxt) (cs, xs) @ lthms_containing ctxt (cs, xs) |> sort_wrt #1; val len = length facts; - val limit = if_none opt_limit (! thms_containing_limit); + val limit = getOpt (opt_limit, ! thms_containing_limit); in if empty_idx andalso not (null ss) then warning "thms_containing: no consts/frees in specification" diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/rule_cases.ML --- a/src/Pure/Isar/rule_cases.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/rule_cases.ML Thu Mar 03 12:43:01 2005 +0100 @@ -73,7 +73,7 @@ fun get thm = let - val n = if_none (lookup_consumes thm) 0; + val n = getOpt (lookup_consumes thm, 0); val ss = (case lookup_case_names thm of NONE => map Library.string_of_int (1 upto (Thm.nprems_of thm - n)) @@ -120,7 +120,7 @@ fun make is_open split (sg, prop) names = let val nprems = Logic.count_prems (prop, 0) in - foldr (fn (name, (cases, i)) => (prep_case is_open split sg prop name i :: cases, i - 1)) + Library.foldr (fn (name, (cases, i)) => (prep_case is_open split sg prop name i :: cases, i - 1)) (Library.drop (length names - nprems, names), ([], length names)) |> #1 end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/session.ML --- a/src/Pure/Isar/session.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/session.ML Thu Mar 03 12:43:01 2005 +0100 @@ -73,7 +73,7 @@ fun get_rpath rpath_str = (if rpath_str = "" then () else - if is_some (!rpath) then + if isSome (!rpath) then error "Path for remote theory browsing information may only be set once" else rpath := SOME (Url.unpack rpath_str); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Isar/toplevel.ML --- a/src/Pure/Isar/toplevel.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Isar/toplevel.ML Thu Mar 03 12:43:01 2005 +0100 @@ -200,7 +200,7 @@ fun interruptible f x = let val y = ref (NONE: node History.T option); - in raise_interrupt (fn () => y := SOME (f x)) (); the (! y) end; + in raise_interrupt (fn () => y := SOME (f x)) (); valOf (! y) end; val rollback = ERROR_MESSAGE "Stale signature encountered after succesful execution!"; @@ -221,7 +221,7 @@ (mk_state (transform_error (interruptible (trans f')) node), NONE) handle exn => (mk_state node, SOME exn); in - if is_stale result then return (mk_state node, SOME (if_none opt_exn rollback)) + if is_stale result then return (mk_state node, SOME (getOpt (opt_exn, rollback))) else return (result, opt_exn) end; @@ -480,7 +480,7 @@ | exn_message (TERM (msg, ts)) = msg_on_debug (TERM (msg, ts)) | exn_message (THM (msg, i, thms)) = msg_on_debug (THM (msg, i, thms)) | exn_message Option = raised "Option" - | exn_message (Library.LIST msg) = raised_msg "Library.LIST" msg + | exn_message UnequalLengths = raised "UnequalLengths" | exn_message exn = General.exnMessage exn and fail_message kind ((name, pos), exn) = "Error in " ^ kind ^ " " ^ quote name ^ Position.str_of pos ^ ":\n" ^ exn_message exn; @@ -502,7 +502,7 @@ (if ! Output.timing andalso not no_timing then (warning (command_msg "" tr); timeap) else I) (apply_trans int trans) state; val _ = if int andalso print andalso not (! quiet) then print_state false result else (); - in (result, apsome (fn UNDEF => type_error tr state | exn => exn) opt_exn) end; + in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_exn) end; in diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Proof/extraction.ML Thu Mar 03 12:43:01 2005 +0100 @@ -86,7 +86,7 @@ fun merge_rules ({next, rs = rs1, net} : rules) ({next = next2, rs = rs2, ...} : rules) = - foldr add_rule (rs2 \\ rs1, {next = next, rs = rs1, net = net}); + Library.foldr add_rule (rs2 \\ rs1, {next = next, rs = rs1, net = net}); fun condrew sign rules procs = let @@ -105,15 +105,15 @@ in get_first (fn (_, (prems, (tm1, tm2))) => let - fun ren t = if_none (Term.rename_abs tm1 tm t) t; + fun ren t = getOpt (Term.rename_abs tm1 tm t, t); val inc = Logic.incr_indexes ([], maxidx_of_term tm + 1); val env as (Tenv, tenv) = Pattern.match tsig (inc tm1, tm); val prems' = map (pairself (subst_vars env o inc o ren)) prems; val env' = Envir.Envir - {maxidx = foldl Int.max + {maxidx = Library.foldl Int.max (~1, map (Int.max o pairself maxidx_of_term) prems'), iTs = Vartab.make Tenv, asol = Vartab.make tenv}; - val env'' = foldl (fn (env, p) => + val env'' = Library.foldl (fn (env, p) => Pattern.unify (sign, env, [pairself (lookup rew) p])) (env', prems') in SOME (Envir.norm_term env'' (inc (ren tm2))) end handle Pattern.MATCH => NONE | Pattern.Unif => NONE) @@ -147,7 +147,7 @@ let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) in Abst (a, SOME T, prf_abstract_over t prf) end; -val mkabs = foldr (fn (v, t) => Abs ("x", fastype_of v, abstract_over (v, t))); +val mkabs = Library.foldr (fn (v, t) => Abs ("x", fastype_of v, abstract_over (v, t))); fun strip_abs 0 t = t | strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t @@ -156,7 +156,7 @@ fun prf_subst_TVars tye = map_proof_terms (subst_TVars tye) (typ_subst_TVars tye); -fun relevant_vars types prop = foldr (fn +fun relevant_vars types prop = Library.foldr (fn (Var ((a, i), T), vs) => (case strip_type T of (_, Type (s, _)) => if s mem types then a :: vs else vs | _ => vs) @@ -254,7 +254,7 @@ defs, expand, prep} = ExtractionData.get thy; in ExtractionData.put - {realizes_eqns = foldr add_rule (map (prep_eq thy) eqns, realizes_eqns), + {realizes_eqns = Library.foldr add_rule (map (prep_eq thy) eqns, realizes_eqns), typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = expand, prep = prep} thy end @@ -272,7 +272,7 @@ in ExtractionData.put {realizes_eqns = realizes_eqns, realizers = realizers, - typeof_eqns = foldr add_rule (eqns', typeof_eqns), + typeof_eqns = Library.foldr add_rule (eqns', typeof_eqns), types = types, defs = defs, expand = expand, prep = prep} thy end @@ -311,7 +311,7 @@ in ExtractionData.put {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, - realizers = foldr Symtab.update_multi + realizers = Library.foldr Symtab.update_multi (map (prep_rlz thy) (rev rs), realizers), defs = defs, expand = expand, prep = prep} thy end @@ -320,7 +320,7 @@ let val {realizes_eqns, typeof_eqns, defs, types, ...} = ExtractionData.get thy; - val procs = flat (map (fst o snd) types); + val procs = List.concat (map (fst o snd) types); val rtypes = map fst types; val eqns = Net.merge (#net realizes_eqns, #net typeof_eqns, K false); val thy' = add_syntax thy; @@ -344,7 +344,7 @@ (procs @ [typeof_proc (Sign.defaultS sign) vs, rlz_proc])) (Const ("realizes", T --> propT --> propT) $ (if T = nullT then t else list_comb (t, vars')) $ prop); - val r = foldr forall_intr (map (get_var_type r') vars, r'); + val r = Library.foldr forall_intr (map (get_var_type r') vars, r'); val prf = Reconstruct.reconstruct_proof sign r (rd s2); in (name, (vs, (t, prf))) end end; @@ -359,7 +359,7 @@ val sign = sign_of thy'; val {realizes_eqns, typeof_eqns, defs, types, ...} = ExtractionData.get thy'; - val procs = flat (map (fst o snd) types); + val procs = List.concat (map (fst o snd) types); val eqns = Net.merge (#net realizes_eqns, #net typeof_eqns, K false); val prop' = Pattern.rewrite_term (Sign.tsig_of sign) (map (Logic.dest_equals o prop_of) defs) [] prop; @@ -382,7 +382,7 @@ (case strip_comb (fst (Logic.dest_equals (prop_of thm))) of (Const _, ts) => forall is_Var ts andalso null (duplicates ts) andalso exists (fn thy => - is_some (Symtab.lookup (#axioms (rep_theory thy), name))) + isSome (Symtab.lookup (#axioms (rep_theory thy), name))) (thy :: ancestors_of thy) | _ => false) handle TERM _ => false; @@ -402,7 +402,7 @@ expand = (name, prop_of thm) ins expand, prep = prep}) thy, thm) end; -fun add_expand_thms thms thy = foldl (fst o add_expand_thm) (thy, thms); +fun add_expand_thms thms thy = Library.foldl (fst o add_expand_thm) (thy, thms); (** types with computational content **) @@ -427,10 +427,10 @@ val tsg = Sign.tsig_of sg; val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; - val procs = flat (map (fst o snd) types); + val procs = List.concat (map (fst o snd) types); val rtypes = map fst types; val typroc = typeof_proc (Sign.defaultS sg); - val prep = if_none prep (K I) sg o ProofRewriteRules.elim_defs sg false defs o + val prep = getOpt (prep, K I) sg o ProofRewriteRules.elim_defs sg false defs o Reconstruct.expand_proof sg (("", NONE) :: map (apsnd SOME) expand); val rrews = Net.merge (#net realizes_eqns, #net typeof_eqns, K false); @@ -448,10 +448,10 @@ end else (vs', tye) - in foldr add_args (take (n, vars) ~~ take (n, ts), ([], [])) end; + in Library.foldr add_args (Library.take (n, vars) ~~ Library.take (n, ts), ([], [])) end; - fun find vs = apsome snd o find_first (curry eq_set vs o fst); - fun find' s = map snd o filter (equal s o fst) + fun find vs = Option.map snd o find_first (curry eq_set vs o fst); + fun find' s = map snd o List.filter (equal s o fst) fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw (condrew sg rrews (procs @ [typroc vs, rlz_proc])) (list_abs @@ -543,7 +543,7 @@ val (defs'', corr_prf) = corr (d + 1) defs' vs' [] [] [] prf' prf' NONE; val corr_prop = Reconstruct.prop_of corr_prf; - val corr_prf' = foldr forall_intr_prf + val corr_prf' = Library.foldr forall_intr_prf (map (get_var_type corr_prop) (vfs_of prop), proof_combt (PThm ((corr_name name vs', []), corr_prf, corr_prop, SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop)) @@ -629,7 +629,7 @@ val nt = Envir.beta_norm t; val args = filter_out (fn v => tname_of (body_type (fastype_of v)) mem rtypes) (vfs_of prop); - val args' = filter (fn v => Logic.occs (v, nt)) args; + val args' = List.filter (fn v => Logic.occs (v, nt)) args; val t' = mkabs (args', nt); val T = fastype_of t'; val cname = extr_name s vs'; @@ -651,7 +651,7 @@ PAxm (cname ^ "_def", eqn, SOME (map TVar (term_tvars eqn))))) %% corr_prf; val corr_prop = Reconstruct.prop_of corr_prf'; - val corr_prf'' = foldr forall_intr_prf + val corr_prf'' = Library.foldr forall_intr_prf (map (get_var_type corr_prop) (vfs_of prop), proof_combt (PThm ((corr_name s vs', []), corr_prf', corr_prop, SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop)); @@ -690,7 +690,7 @@ quote name ^ " has no computational content") in (Reconstruct.reconstruct_proof sign prop prf, vs) end; - val defs = foldl (fn (defs, (prf, vs)) => + val defs = Library.foldl (fn (defs, (prf, vs)) => fst (extr 0 defs vs [] [] [] prf)) ([], map prep_thm thms); val {path, ...} = Sign.rep_sg sg; @@ -716,8 +716,8 @@ in thy |> Theory.absolute_path |> - curry (foldr add_def) defs |> - Theory.add_path (NameSpace.pack (if_none path [])) + curry (Library.foldr add_def) defs |> + Theory.add_path (NameSpace.pack (getOpt (path,[]))) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Proof/proof_rewrite_rules.ML --- a/src/Pure/Proof/proof_rewrite_rules.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Proof/proof_rewrite_rules.ML Thu Mar 03 12:43:01 2005 +0100 @@ -197,7 +197,7 @@ fun strip [] t = t | strip (_ :: xs) (Abs (_, _, t)) = strip xs t; in - strip Ts (foldl (uncurry lambda o Library.swap) (t', frees)) + strip Ts (Library.foldl (uncurry lambda o Library.swap) (t', frees)) end; fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2 @@ -227,9 +227,9 @@ val vs = vars_of prop; val tvars = term_tvars prop; val (_, rhs) = Logic.dest_equals prop; - val rhs' = foldl betapply (subst_TVars (map fst tvars ~~ Ts) - (foldr (fn p => Abs ("", dummyT, abstract_over p)) (vs, rhs)), - map the ts); + val rhs' = Library.foldl betapply (subst_TVars (map fst tvars ~~ Ts) + (Library.foldr (fn p => Abs ("", dummyT, abstract_over p)) (vs, rhs)), + map valOf ts); in change_type (SOME [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs' end @@ -245,7 +245,7 @@ val f = if not r then I else let val cnames = map (fst o dest_Const o fst) defs'; - val thms = flat (map (fn (s, ps) => + val thms = List.concat (map (fn (s, ps) => if s mem defnames then [] else map (pair s o SOME o fst) (filter_out (fn (p, _) => null (term_consts p inter cnames)) ps)) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Proof/proof_syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -92,12 +92,12 @@ fun disambiguate_names thy prf = let val thms = thms_of_proof Symtab.empty prf; - val thms' = map (apsnd (#prop o rep_thm)) (flat + val thms' = map (apsnd (#prop o rep_thm)) (List.concat (map PureThy.thms_of (thy :: Theory.ancestors_of thy))); val tab = Symtab.foldl (fn (tab, (key, ps)) => - let val prop = if_none (assoc (thms', key)) (Bound 0) - in fst (foldr (fn ((prop', prf), x as (tab, i)) => + let val prop = getOpt (assoc (thms', key), Bound 0) + in fst (Library.foldr (fn ((prop', prf), x as (tab, i)) => if prop <> prop' then (Symtab.update ((key ^ "_" ^ string_of_int i, prf), tab), i+1) else x) (ps, (tab, 1))) @@ -109,8 +109,8 @@ | rename (prf % t) = rename prf % t | rename (prf' as PThm ((s, tags), prf, prop, Ts)) = let - val prop' = if_none (assoc (thms', s)) (Bound 0); - val ps = map fst (the (Symtab.lookup (thms, s))) \ prop' + val prop' = getOpt (assoc (thms', s), Bound 0); + val ps = map fst (valOf (Symtab.lookup (thms, s))) \ prop' in if prop = prop' then prf' else PThm ((s ^ "_" ^ string_of_int (length ps - find_index_eq prop ps), tags), prf, prop, Ts) @@ -125,8 +125,8 @@ fun proof_of_term thy tab ty = let val thys = thy :: Theory.ancestors_of thy; - val thms = flat (map thms_of thys); - val axms = flat (map (Symtab.dest o #axioms o rep_theory) thys); + val thms = List.concat (map thms_of thys); + val axms = List.concat (map (Symtab.dest o #axioms o rep_theory) thys); fun mk_term t = (if ty then I else map_term_types (K dummyT)) (Term.no_dummy_patterns t); @@ -180,7 +180,7 @@ val Oraclet = Const ("Oracle", propT --> proofT); val MinProoft = Const ("MinProof", proofT); -val mk_tyapp = foldl (fn (prf, T) => Const ("Appt", +val mk_tyapp = Library.foldl (fn (prf, T) => Const ("Appt", [proofT, itselfT T] ---> proofT) $ prf $ Logic.mk_type T); fun term_of _ (PThm ((name, _), _, _, NONE)) = @@ -192,17 +192,17 @@ mk_tyapp (Const (add_prefix "axm" name, proofT), Ts) | term_of _ (PBound i) = Bound i | term_of Ts (Abst (s, opT, prf)) = - let val T = if_none opT dummyT + let val T = getOpt (opT,dummyT) in Const ("Abst", (T --> proofT) --> proofT) $ Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf)) end | term_of Ts (AbsP (s, t, prf)) = - AbsPt $ if_none t (Const ("dummy_pattern", propT)) $ + AbsPt $ getOpt (t, Const ("dummy_pattern", propT)) $ Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf)) | term_of Ts (prf1 %% prf2) = AppPt $ term_of Ts prf1 $ term_of Ts prf2 | term_of Ts (prf % opt) = - let val t = if_none opt (Const ("dummy_pattern", dummyT)) + let val t = getOpt (opt, Const ("dummy_pattern", dummyT)) in Const ("Appt", [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $ term_of Ts prf $ t @@ -217,9 +217,9 @@ let val (prf', tab) = disambiguate_names thy prf; val thys = thy :: Theory.ancestors_of thy; - val thm_names = filter_out (equal "") (map fst (flat (map thms_of thys))) @ + val thm_names = filter_out (equal "") (map fst (List.concat (map thms_of thys))) @ map fst (Symtab.dest tab); - val axm_names = map fst (flat (map (Symtab.dest o #axioms o rep_theory) thys)); + val axm_names = map fst (List.concat (map (Symtab.dest o #axioms o rep_theory) thys)); val sg = sign_of thy |> add_proof_syntax |> add_proof_atom_consts @@ -232,8 +232,8 @@ fun read_term thy = let val thys = thy :: Theory.ancestors_of thy; - val thm_names = filter_out (equal "") (map fst (flat (map thms_of thys))); - val axm_names = map fst (flat (map (Symtab.dest o #axioms o rep_theory) thys)); + val thm_names = filter_out (equal "") (map fst (List.concat (map thms_of thys))); + val axm_names = map fst (List.concat (map (Symtab.dest o #axioms o rep_theory) thys)); val sg = sign_of thy |> add_proof_syntax |> add_proof_atom_consts diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Proof/proofchecker.ML --- a/src/Pure/Proof/proofchecker.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Proof/proofchecker.ML Thu Mar 03 12:43:01 2005 +0100 @@ -19,8 +19,8 @@ (***** construct a theorem out of a proof term *****) fun lookup_thm thy = - let val tab = foldr Symtab.update - (flat (map thms_of (thy :: Theory.ancestors_of thy)), Symtab.empty) + let val tab = Library.foldr Symtab.update + (List.concat (map thms_of (thy :: Theory.ancestors_of thy)), Symtab.empty) in (fn s => case Symtab.lookup (tab, s) of NONE => error ("Unknown theorem " ^ quote s) @@ -58,7 +58,7 @@ | thm_of _ _ (PAxm (name, _, SOME Ts)) = thm_of_atom (get_axiom thy name) Ts - | thm_of _ Hs (PBound i) = nth_elem (i, Hs) + | thm_of _ Hs (PBound i) = List.nth (Hs, i) | thm_of vs Hs (Abst (s, SOME T, prf)) = let diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Proof/reconstruct.ML --- a/src/Pure/Proof/reconstruct.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Proof/reconstruct.ML Thu Mar 03 12:43:01 2005 +0100 @@ -30,14 +30,14 @@ let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) in all T $ Abs (a, T, abstract_over (t, prop)) end; -fun forall_intr_vfs prop = foldr forall_intr +fun forall_intr_vfs prop = Library.foldr forall_intr (vars_of prop @ sort (make_ord atless) (term_frees prop), prop); fun forall_intr_prf (t, prf) = let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) in Abst (a, SOME T, prf_abstract_over t prf) end; -fun forall_intr_vfs_prf prop prf = foldr forall_intr_prf +fun forall_intr_vfs_prf prop prf = Library.foldr forall_intr_prf (vars_of prop @ sort (make_ord atless) (term_frees prop), prf); fun merge_envs (Envir.Envir {asol=asol1, iTs=iTs1, maxidx=maxidx1}) @@ -57,7 +57,7 @@ (Envir.Envir {iTs = iTs, asol = asol, maxidx = maxidx+1}, TVar (("'t", maxidx+1), s)); -fun mk_abs Ts t = foldl (fn (u, T) => Abs ("", T, u)) (t, Ts); +fun mk_abs Ts t = Library.foldl (fn (u, T) => Abs ("", T, u)) (t, Ts); fun unifyT sg env T U = let @@ -103,14 +103,14 @@ let val (env3, V) = mk_tvar (env2, []) in (t' $ u', V, vTs2, unifyT sg env3 T (U --> V)) end) end - | infer_type sg env Ts vTs (t as Bound i) = (t, nth_elem (i, Ts), vTs, env); + | infer_type sg env Ts vTs (t as Bound i) = (t, List.nth (Ts, i), vTs, env); fun cantunify sg (t, u) = error ("Non-unifiable terms:\n" ^ Sign.string_of_term sg t ^ "\n\n" ^ Sign.string_of_term sg u); fun decompose sg Ts (env, p as (t, u)) = let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify sg p - else apsnd flat (foldl_map (decompose sg Ts) (uT env T U, ts ~~ us)) + else apsnd List.concat (foldl_map (decompose sg Ts) (uT env T U, ts ~~ us)) in case pairself (strip_comb o Envir.head_norm env) p of ((Const c, ts), (Const d, us)) => rigrig c d (unifyT sg) ts us | ((Free c, ts), (Free d, us)) => rigrig c d (unifyT sg) ts us @@ -149,7 +149,7 @@ NONE => foldl_map mk_tvar (env, map snd tvars @ map snd tfrees) | SOME Ts => (env, Ts)); val prop'' = subst_TVars (map fst tvars @ map snd fmap ~~ Ts) - (forall_intr_vfs prop') handle LIST _ => + (forall_intr_vfs prop') handle UnequalLengths => error ("Wrong number of type arguments for " ^ quote (fst (get_name_tags [] prop prf))) in (prop'', change_type (SOME Ts) prf, [], env', vTs) end; @@ -157,7 +157,7 @@ fun head_norm (prop, prf, cnstrts, env, vTs) = (Envir.head_norm env prop, prf, cnstrts, env, vTs); - fun mk_cnstrts env _ Hs vTs (PBound i) = (nth_elem (i, Hs), PBound i, [], env, vTs) + fun mk_cnstrts env _ Hs vTs (PBound i) = (List.nth (Hs, i), PBound i, [], env, vTs) | mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) = let val (env', T) = (case opT of @@ -314,7 +314,7 @@ val head_norm = Envir.head_norm (Envir.empty 0); -fun prop_of0 Hs (PBound i) = nth_elem (i, Hs) +fun prop_of0 Hs (PBound i) = List.nth (Hs, i) | prop_of0 Hs (Abst (s, SOME T, prf)) = all T $ (Abs (s, T, prop_of0 Hs prf)) | prop_of0 Hs (AbsP (s, SOME t, prf)) = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/ast.ML --- a/src/Pure/Syntax/ast.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/ast.ML Thu Mar 03 12:43:01 2005 +0100 @@ -119,7 +119,7 @@ let fun vars_of (Constant _) = [] | vars_of (Variable x) = [x] - | vars_of (Appl asts) = flat (map vars_of asts); + | vars_of (Appl asts) = List.concat (map vars_of asts); fun unique (x :: xs) = not (x mem xs) andalso unique xs | unique [] = true; @@ -142,7 +142,7 @@ | fold_ast _ [y] = y | fold_ast c (x :: xs) = Appl [Constant c, x, fold_ast c xs]; -fun fold_ast_p c = foldr (fn (x, xs) => Appl [Constant c, x, xs]); +fun fold_ast_p c = Library.foldr (fn (x, xs) => Appl [Constant c, x, xs]); fun fold_ast2 _ _ [] = raise Match | fold_ast2 _ c1 [y] = Appl [Constant c1, y] @@ -190,7 +190,7 @@ (case (ast, pat) of (Appl asts, Appl pats) => let val a = length asts and p = length pats in - if a > p then (Appl (take (p, asts)), drop (p, asts)) + if a > p then (Appl (Library.take (p, asts)), Library.drop (p, asts)) else (ast, []) end | _ => (ast, [])); @@ -210,7 +210,7 @@ val changes = ref 0; fun subst _ (ast as Constant _) = ast - | subst env (Variable x) = the (Symtab.lookup (env, x)) + | subst env (Variable x) = valOf (Symtab.lookup (env, x)) | subst env (Appl asts) = Appl (map (subst env) asts); fun try_rules ((lhs, rhs) :: pats) ast = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/lexicon.ML --- a/src/Pure/Syntax/lexicon.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/lexicon.ML Thu Mar 03 12:43:01 2005 +0100 @@ -195,7 +195,7 @@ let fun assoc [] = [] | assoc ((keyi, xi) :: pairs) = - if is_none keyi orelse matching_tokens (the keyi, key) then + if is_none keyi orelse matching_tokens (valOf keyi, key) then assoc pairs @ xi else assoc pairs; in assoc list end; @@ -292,7 +292,7 @@ Scan.one Symbol.is_blank >> K NONE; in (case Scan.error (Scan.finite Symbol.stopper (Scan.repeat scan_token)) chs of - (toks, []) => mapfilter I toks @ [EndToken] + (toks, []) => List.mapPartial I toks @ [EndToken] | (_, cs) => error ("Inner lexical error at: " ^ quote (implode cs))) end; @@ -345,7 +345,7 @@ val scan = $$ "?" |-- scan_indexname >> var || Scan.any Symbol.not_eof >> (free o implode); - in the (Scan.read Symbol.stopper scan (Symbol.explode str)) end; + in valOf (Scan.read Symbol.stopper scan (Symbol.explode str)) end; (* variable kinds *) @@ -360,7 +360,7 @@ (* read_nat *) fun read_nat str = - apsome (#1 o Library.read_int) (Scan.read Symbol.stopper scan_digits1 (Symbol.explode str)); + Option.map (#1 o Library.read_int) (Scan.read Symbol.stopper scan_digits1 (Symbol.explode str)); (* read_xnum *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/mixfix.ML --- a/src/Pure/Syntax/mixfix.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/mixfix.ML Thu Mar 03 12:43:01 2005 +0100 @@ -151,7 +151,7 @@ | _ => error ("Bad mixfix declaration for type: " ^ quote t)) end; - val mfix = mapfilter mfix_of type_decls; + val mfix = List.mapPartial mfix_of type_decls; val xconsts = map name_of type_decls; in SynExt.syn_ext mfix xconsts ([], [], [], []) [] ([], []) end; @@ -189,9 +189,9 @@ fun binder (c, _, Binder (sy, _, _)) = SOME (sy, c) | binder _ = NONE; - val mfix = flat (map mfix_of const_decls); + val mfix = List.concat (map mfix_of const_decls); val xconsts = map name_of const_decls; - val binders = mapfilter binder const_decls; + val binders = List.mapPartial binder const_decls; val binder_trs = map SynTrans.mk_binder_tr binders; val binder_trs' = map (apsnd SynTrans.non_typed_tr' o SynTrans.mk_binder_tr' o swap) binders; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/parser.ML --- a/src/Pure/Syntax/parser.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/parser.ML Thu Mar 03 12:43:01 2005 +0100 @@ -88,7 +88,7 @@ val (added_starts, lambdas') = if is_none new_chain then ([], lambdas) else let (*lookahead of chain's source*) - val ((from_nts, from_tks), _) = Array.sub (prods, the new_chain); + val ((from_nts, from_tks), _) = Array.sub (prods, valOf new_chain); (*copy from's lookahead to chain's destinations*) fun copy_lookahead [] added = added @@ -152,8 +152,8 @@ let val new_lambda = is_none tk andalso nts subset lambdas; - val new_tks = (if is_some tk then [the tk] else []) @ - foldl token_union ([], map starts_for_nt nts) \\ + val new_tks = (if isSome tk then [valOf tk] else []) @ + Library.foldl token_union ([], map starts_for_nt nts) \\ l_starts; val added_tks' = token_union (new_tks, added_tks); @@ -193,7 +193,7 @@ (*existing productions whose lookahead may depend on l*) val tk_prods = assocs nt_prods - (SOME (hd l_starts handle LIST _ => UnknownStart)); + (SOME (hd l_starts handle Empty => UnknownStart)); (*add_lambda is true if an existing production of the nt produces lambda due to the new lambda NT l*) @@ -231,15 +231,15 @@ (*insert production into grammar*) val (added_starts', prod_count') = - if is_some new_chain then (added_starts', prod_count) + if isSome new_chain then (added_starts', prod_count) (*don't store chain production*) else let (*lookahead tokens of new production and on which NTs lookahead depends*) val (start_tk, start_nts) = lookahead_dependency lambdas' rhs []; - val start_tks = foldl token_union - (if is_some start_tk then [the start_tk] else [], + val start_tks = Library.foldl token_union + (if isSome start_tk then [valOf start_tk] else [], map starts_for_nt start_nts); val opt_starts = (if new_lambda then [NONE] @@ -265,8 +265,8 @@ (*store new production*) fun store [] prods is_new = - (prods, if is_some prod_count andalso is_new then - apsome (fn x => x+1) prod_count + (prods, if isSome prod_count andalso is_new then + Option.map (fn x => x+1) prod_count else prod_count, is_new) | store (tk :: tks) prods is_new = let val tk_prods = assocs prods tk; @@ -274,7 +274,7 @@ (*if prod_count = NONE then we can assume that grammar does not contain new production already*) val (tk_prods', is_new') = - if is_some prod_count then + if isSome prod_count then if new_prod mem tk_prods then (tk_prods, false) else (new_prod :: tk_prods, true) else (new_prod :: tk_prods, true); @@ -328,8 +328,8 @@ (*test if production could already be associated with a member of new_tks*) val lambda = length depends > 1 orelse - not (null depends) andalso is_some tk - andalso the tk mem new_tks; + not (null depends) andalso isSome tk + andalso valOf tk mem new_tks; (*associate production with new starting tokens*) fun copy [] nt_prods = nt_prods @@ -399,7 +399,7 @@ fun pretty_symb (Terminal (Token s)) = Pretty.quote (Pretty.str s) | pretty_symb (Terminal tok) = Pretty.str (str_of_token tok) | pretty_symb (Nonterminal (tag, p)) = - let val name = fst (the (find_first (fn (n, t) => t = tag) taglist)); + let val name = fst (valOf (find_first (fn (n, t) => t = tag) taglist)); in Pretty.str (name ^ "[" ^ string_of_int p ^ "]") end; fun pretty_const "" = [] @@ -416,11 +416,11 @@ fun prod_of_chain from = ([Nonterminal (from, ~1)], "", ~1); val nt_prods = - foldl (op union) ([], map snd (snd (Array.sub (prods, tag)))) @ + Library.foldl (op union) ([], map snd (snd (Array.sub (prods, tag)))) @ map prod_of_chain (assocs chains tag); in map (pretty_prod name) nt_prods end; - in flat (map pretty_nt taglist) end; + in List.concat (map pretty_nt taglist) end; (** Operations on gramars **) @@ -537,7 +537,7 @@ | store_tag nt_count tags tag = let val (nt_count', tags', tag') = get_tag nt_count tags - (fst (the (find_first (fn (n, t) => t = tag) tag_list))); + (fst (valOf (find_first (fn (n, t) => t = tag) tag_list))); in Array.update (table, tag, tag'); store_tag nt_count' tags' (tag-1) end; @@ -563,7 +563,7 @@ fun process_nt ~1 result = result | process_nt nt result = let - val nt_prods = foldl (op union) + val nt_prods = Library.foldl (op union) ([], map snd (snd (Array.sub (prods2, nt)))); val lhs_tag = convert_tag nt; @@ -616,11 +616,11 @@ (*Get all rhss with precedence >= minPrec*) -fun getRHS minPrec = filter (fn (_, _, prec:int) => prec >= minPrec); +fun getRHS minPrec = List.filter (fn (_, _, prec:int) => prec >= minPrec); (*Get all rhss with precedence >= minPrec and < maxPrec*) fun getRHS' minPrec maxPrec = - filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec); + List.filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec); (*Make states using a list of rhss*) fun mkStates i minPrec lhsID rhss = @@ -656,19 +656,19 @@ in update (used, []) end; fun getS A maxPrec Si = - filter + List.filter (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= maxPrec | _ => false) Si; fun getS' A maxPrec minPrec Si = - filter + List.filter (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec > minPrec andalso prec <= maxPrec | _ => false) Si; fun getStates Estate i ii A maxPrec = - filter + List.filter (fn (_, _, _, Nonterminal (B, prec) :: _, _, _) => A = B andalso prec <= maxPrec | _ => false) @@ -701,7 +701,7 @@ fun token_assoc2 (list, key) = let fun assoc [] result = result | assoc ((keyi, pi) :: pairs) result = - if is_some keyi andalso matching_tokens (the keyi, key) + if isSome keyi andalso matching_tokens (valOf keyi, key) orelse include_none andalso is_none keyi then assoc pairs (pi @ result) else assoc pairs result; @@ -841,12 +841,12 @@ end; val nts = - mapfilter (fn (_, _, _, Nonterminal (a, prec) :: _, _, _) => + List.mapPartial (fn (_, _, _, Nonterminal (a, prec) :: _, _, _) => SOME (a, prec) | _ => NONE) (Array.sub (stateset, i-1)); val allowed = distinct (get_starts nts [] @ - (mapfilter (fn (_, _, _, Terminal a :: _, _, _) => SOME a + (List.mapPartial (fn (_, _, _, Terminal a :: _, _, _) => SOME a | _ => NONE) (Array.sub (stateset, i-1)))); in syntax_error (if prev_token = EndToken then indata @@ -863,7 +863,7 @@ end)); -fun get_trees l = mapfilter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) +fun get_trees l = List.mapPartial (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) l; fun earley prods tags chains startsymbol indata = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/printer.ML --- a/src/Pure/Syntax/printer.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/printer.ML Thu Mar 03 12:43:01 2005 +0100 @@ -188,11 +188,11 @@ (*find tab for mode*) fun get_tab prtabs mode = - if_none (assoc (prtabs, mode)) Symtab.empty; + getOpt (assoc (prtabs, mode), Symtab.empty); (*collect tabs for mode hierarchy (default "")*) fun tabs_of prtabs modes = - mapfilter (fn mode => assoc (prtabs, mode)) (modes @ [""]); + List.mapPartial (fn mode => assoc (prtabs, mode)) (modes @ [""]); (* xprod_to_fmt *) @@ -244,9 +244,9 @@ fun extend_prtabs prtabs mode xprods = let - val fmts = mapfilter xprod_to_fmt xprods; + val fmts = List.mapPartial xprod_to_fmt xprods; val tab = get_tab prtabs mode; - val new_tab = foldr Symtab.update_multi (rev fmts, tab); + val new_tab = Library.foldr Symtab.update_multi (rev fmts, tab); in overwrite (prtabs, (mode, new_tab)) end; fun merge_prtabs prtabs1 prtabs2 = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/syn_ext.ML --- a/src/Pure/Syntax/syn_ext.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/syn_ext.ML Thu Mar 03 12:43:01 2005 +0100 @@ -144,9 +144,9 @@ | del_of _ = NONE; fun dels_of (XProd (_, xsymbs, _, _)) = - mapfilter del_of xsymbs; + List.mapPartial del_of xsymbs; in - map Symbol.explode (distinct (flat (map dels_of xprods))) + map Symbol.explode (distinct (List.concat (map dels_of xprods))) end; @@ -206,14 +206,14 @@ $$ "'" -- Scan.one Symbol.is_blank >> K NONE; val scan_symbs = Scan.repeat scan_symb --| Scan.ahead (Scan.one (not_equal "'")); - val read_symbs = mapfilter I o the o Scan.read Symbol.stopper scan_symbs; + val read_symbs = List.mapPartial I o valOf o Scan.read Symbol.stopper scan_symbs; fun unique_index xsymbs = - if length (filter is_index xsymbs) <= 1 then xsymbs + if length (List.filter is_index xsymbs) <= 1 then xsymbs else error "Duplicate index arguments (\\)"; in val read_mixfix = unique_index o read_symbs o Symbol.explode; - val mfix_args = length o filter is_argument o read_mixfix; + val mfix_args = length o List.filter is_argument o read_mixfix; val escape_mfix = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode; end; @@ -261,7 +261,7 @@ val raw_symbs = read_mixfix sy handle ERROR => err_in_mfix "" mfix; - val args = filter (fn Argument _ => true | _ => false) raw_symbs; + val args = List.filter (fn Argument _ => true | _ => false) raw_symbs; val (const', typ', parse_rules) = if not (exists is_index args) then (const, typ, []) else @@ -293,11 +293,11 @@ val symbs' = map (logify_types copy_prod) symbs; val xprod = XProd (lhs', symbs', const', pri); - val _ = (seq check_pri pris; check_pri pri; check_blocks symbs'); + val _ = (List.app check_pri pris; check_pri pri; check_blocks symbs'); val xprod' = if Lexicon.is_terminal lhs' then err_in_mfix ("Illegal lhs: " ^ lhs') mfix else if const <> "" then xprod - else if length (filter is_argument symbs') <> 1 then + else if length (List.filter is_argument symbs') <> 1 then err_in_mfix "Copy production must have exactly one argument" mfix else if exists is_terminal symbs' then xprod else XProd (lhs', map rem_pri symbs', "", chain_pri); @@ -330,7 +330,7 @@ print_ast_translation) = trfuns; val (xprods, parse_rules') = map (mfix_to_xprod convert is_logtype) mfixes - |> split_list |> apsnd (rev o flat); + |> split_list |> apsnd (rev o List.concat); val mfix_consts = distinct (map (fn Mfix x => #3 x) mfixes @ map (fn XProd x => #3 x) xprods); in SynExt { diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/syn_trans.ML --- a/src/Pure/Syntax/syn_trans.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/syn_trans.ML Thu Mar 03 12:43:01 2005 +0100 @@ -202,7 +202,7 @@ (* implicit structures *) fun the_struct structs i = - if 1 <= i andalso i <= length structs then Library.nth_elem (i - 1, structs) + if 1 <= i andalso i <= length structs then List.nth (structs, i - 1) else raise ERROR_MESSAGE ("Illegal reference to implicit structure #" ^ string_of_int i); fun struct_tr structs (*"_struct"*) [Const ("_indexdefault", _)] = @@ -272,7 +272,7 @@ fun abs_tr' tm = - foldr (fn (x, t) => Lexicon.const "_abs" $ x $ t) + Library.foldr (fn (x, t) => Lexicon.const "_abs" $ x $ t) (strip_abss strip_abs_vars strip_abs_body (eta_contr tm)); fun atomic_abs_tr' (x, T, t) = @@ -481,8 +481,8 @@ | ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok); val exn_results = map (capture ast_of) pts; - val exns = mapfilter get_exn exn_results; - val results = mapfilter get_result exn_results + val exns = List.mapPartial get_exn exn_results; + val results = List.mapPartial get_result exn_results in (case (results, exns) of ([], TRANSLATION (a, exn) :: _) => @@ -510,8 +510,8 @@ | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]); val exn_results = map (capture term_of) asts; - val exns = mapfilter get_exn exn_results; - val results = mapfilter get_result exn_results + val exns = List.mapPartial get_exn exn_results; + val results = List.mapPartial get_result exn_results in (case (results, exns) of ([], TRANSLATION (a, exn) :: _) => diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Syntax/syntax.ML --- a/src/Pure/Syntax/syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Syntax/syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -84,7 +84,7 @@ (* parse (ast) translations *) -fun lookup_tr tab c = apsome fst (Symtab.lookup (tab, c)); +fun lookup_tr tab c = Option.map fst (Symtab.lookup (tab, c)); fun err_dup_trfuns name cs = error ("More than one " ^ name ^ " for " ^ commas_quote cs); @@ -101,7 +101,7 @@ (* print (ast) translations *) fun lookup_tr' tab c = map fst (Symtab.lookup_multi (tab, c)); -fun extend_tr'tab tab trfuns = foldr Symtab.update_multi (map mk_trfun trfuns, tab); +fun extend_tr'tab tab trfuns = Library.foldr Symtab.update_multi (map mk_trfun trfuns, tab); fun merge_tr'tabs tab1 tab2 = Symtab.merge_multi' eq_trfuns (tab1, tab2); @@ -109,8 +109,8 @@ (** tables of token translation functions **) fun lookup_tokentr tabs modes = - let val trs = gen_distinct eq_fst (flat (map (assocs tabs) (modes @ [""]))) - in fn c => apsome fst (assoc (trs, c)) end; + let val trs = gen_distinct eq_fst (List.concat (map (assocs tabs) (modes @ [""]))) + in fn c => Option.map fst (assoc (trs, c)) end; fun merge_tokentrtabs tabs1 tabs2 = let @@ -138,7 +138,7 @@ fun ins_tokentr (ts, (m, c, f)) = overwrite (ts, (m, ("_" ^ c, (f, stamp ())) :: assocs ts m)); in - merge_tokentrtabs tabs (foldl ins_tokentr ([], tokentrs)) + merge_tokentrtabs tabs (Library.foldl ins_tokentr ([], tokentrs)) end; @@ -147,14 +147,14 @@ type ruletab = (Ast.ast * Ast.ast) list Symtab.table; -fun dest_ruletab tab = flat (map snd (Symtab.dest tab)); +fun dest_ruletab tab = List.concat (map snd (Symtab.dest tab)); fun lookup_ruletab tab a = Symtab.lookup_multi (tab, a); (* empty, extend, merge ruletabs *) fun extend_ruletab tab rules = - foldr Symtab.update_multi (map (fn r => (Ast.head_of_rule r, r)) rules, tab); + Library.foldr Symtab.update_multi (map (fn r => (Ast.head_of_rule r, r)) rules, tab); fun merge_ruletabs tab1 tab2 = Symtab.merge_multi' (op =) (tab1, tab2); @@ -343,7 +343,7 @@ else (warning ("Ambiguous input " ^ quote str); warning "produces the following parse trees:"; - seq show_pt pts) + List.app show_pt pts) else (); SynTrans.pts_to_asts (lookup_tr parse_ast_trtab) pts end; @@ -426,8 +426,8 @@ fun prep_rules rd_pat raw_rules = let val rules = map (map_trrule rd_pat) raw_rules in - (map check_rule (mapfilter parse_rule rules), - map check_rule (mapfilter print_rule rules)) + (map check_rule (List.mapPartial parse_rule rules), + map check_rule (List.mapPartial print_rule rules)) end diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Thy/present.ML --- a/src/Pure/Thy/present.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Thy/present.ML Thu Mar 03 12:43:01 2005 +0100 @@ -115,8 +115,8 @@ {name = name, ID = ID_of session name, dir = sess_name, path = if null session then "" else - if is_some remote_path andalso not is_local then - Url.pack (Url.append (the remote_path) (Url.File + if isSome remote_path andalso not is_local then + Url.pack (Url.append (valOf remote_path) (Url.File (Path.append (Path.make session) (html_path name)))) else Path.pack (Path.append (mk_rel_path curr_sess session) (html_path name)), unfold = false, @@ -296,7 +296,7 @@ val parent_index_path = Path.append Path.parent index_path; val index_up_lnk = if first_time then - Url.append (the remote_path) (Url.File (Path.append sess_prefix parent_index_path)) + Url.append (valOf remote_path) (Url.File (Path.append sess_prefix parent_index_path)) else Url.File parent_index_path; val readme = if File.exists readme_html_path then SOME readme_html_path @@ -304,8 +304,8 @@ else NONE; val index_text = HTML.begin_index (index_up_lnk, parent_name) - (Url.File index_path, session_name) (apsome Url.File readme) - (apsome Url.File document_path) (Url.unpack "medium.html"); + (Url.File index_path, session_name) (Option.map Url.File readme) + (Option.map Url.File document_path) (Url.unpack "medium.html"); in session_info := SOME (make_session_info (name, parent_name, session_name, path, html_prefix, info, doc, doc_graph, doc_prefix1, doc_prefix2, remote_path, verbose, readme)); @@ -363,7 +363,7 @@ Buffer.write (Path.append html_prefix (html_path a)) (Buffer.add HTML.end_theory html); val opt_graphs = - if doc_graph andalso (is_some doc_prefix1 orelse is_some doc_prefix2) then + if doc_graph andalso (isSome doc_prefix1 orelse isSome doc_prefix2) then SOME (isatool_browser graph) else NONE; @@ -374,7 +374,7 @@ (File.write (Path.append path graph_pdf_path) pdf; File.write (Path.append path graph_eps_path) eps)); write_tex_index tex_index path; - seq (finish_tex path) thys); + List.app (finish_tex path) thys); in conditional info (fn () => (File.mkdir (Path.append html_prefix session_path); @@ -385,11 +385,11 @@ (case readme of NONE => () | SOME path => File.copy path (Path.append html_prefix path)); write_graph graph (Path.append html_prefix graph_path); copy_files (Path.unpack "~~/lib/browser/GraphBrowser.jar") html_prefix; - seq (fn (a, txt) => File.write (Path.append html_prefix (Path.basic a)) txt) + List.app (fn (a, txt) => File.write (Path.append html_prefix (Path.basic a)) txt) (HTML.applet_pages name (Url.File index_path, name)); copy_files (Path.unpack "~~/lib/html/isabelle.css") html_prefix; - seq finish_html thys; - seq (uncurry File.write) files; + List.app finish_html thys; + List.app (uncurry File.write) files; conditional verbose (fn () => std_error ("Browser info at " ^ show_path html_prefix ^ "\n")))); @@ -426,8 +426,8 @@ fun parent_link remote_path curr_session name = let val {name = _, session, is_local} = get_info (ThyInfo.theory name) in (if null session then NONE else - SOME (if is_some remote_path andalso not is_local then - Url.append (the remote_path) (Url.File + SOME (if isSome remote_path andalso not is_local then + Url.append (valOf remote_path) (Url.File (Path.append (Path.make session) (html_path name))) else Url.File (Path.append (mk_rel_path curr_session session) (html_path name))), name) @@ -439,7 +439,7 @@ val parents = map (parent_link remote_path path) raw_parents; val ml_path = ThyLoad.ml_path name; val files = map (apsnd SOME) orig_files @ - (if is_some (ThyLoad.check_file optpath ml_path) then [(ml_path, NONE)] else []); + (if isSome (ThyLoad.check_file optpath ml_path) then [(ml_path, NONE)] else []); fun prep_file (raw_path, loadit) = (case ThyLoad.check_file optpath raw_path of @@ -473,7 +473,7 @@ add_graph_entry entry; add_html_index (HTML.theory_entry (Url.File (html_path name), name)); add_tex_index (Latex.theory_entry name); - BrowserInfoData.put {name = sess_name, session = path, is_local = is_some remote_path} thy + BrowserInfoData.put {name = sess_name, session = path, is_local = isSome remote_path} thy end); @@ -481,7 +481,7 @@ fun add_hook f = hooks := (f :: ! hooks); fun results k xs = - (seq (fn f => (try (fn () => f k xs) (); ())) (! hooks); + (List.app (fn f => (try (fn () => f k xs) (); ())) (! hooks); with_session () (fn _ => with_context add_html (HTML.results k xs))); fun theorem a th = results "theorem" [(a, [th])]; @@ -524,7 +524,7 @@ val known_syms = known "syms.lst"; val known_ctrls = known "ctrls.lst"; - val _ = srcs |> seq (fn (name, base, txt) => + val _ = srcs |> List.app (fn (name, base, txt) => Symbol.explode txt |> Latex.symbol_source (known_syms, known_ctrls) (Path.pack base) |> File.write (Path.append doc_path (tex_path name))); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Thy/thm_deps.ML --- a/src/Pure/Thy/thm_deps.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Thy/thm_deps.ML Thu Mar 03 12:43:01 2005 +0100 @@ -34,7 +34,7 @@ | make_deps_graph (p, prf1 %% prf2) = make_deps_graph (make_deps_graph (p, prf1), prf2) | make_deps_graph (p, prf % _) = make_deps_graph (p, prf) - | make_deps_graph (p, MinProof prfs) = foldl make_deps_graph (p, prfs) + | make_deps_graph (p, MinProof prfs) = Library.foldl make_deps_graph (p, prfs) | make_deps_graph (p as (gra, parents), prf) = let val ((name, tags), prf') = dest_thm_axm prf in @@ -68,7 +68,7 @@ fun thm_deps thms = let val _ = writeln "Generating graph ..."; - val gra = map snd (Symtab.dest (fst (foldl make_deps_graph ((Symtab.empty, []), + val gra = map snd (Symtab.dest (fst (Library.foldl make_deps_graph ((Symtab.empty, []), map Thm.proof_of thms)))); val path = File.tmp_path (Path.unpack "theorems.graph"); val _ = Present.write_graph gra path; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Thy/thy_info.ML --- a/src/Pure/Thy/thy_info.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Thy/thy_info.ML Thu Mar 03 12:43:01 2005 +0100 @@ -62,7 +62,7 @@ val hooks = ref ([]: (action -> string -> unit) list); in fun add_hook f = hooks := f :: ! hooks; - fun perform action name = seq (fn f => (try (fn () => f action name) (); ())) (! hooks); + fun perform action name = List.app (fn f => (try (fn () => f action name) (); ())) (! hooks); end; @@ -86,7 +86,7 @@ handle Graph.CYCLES namess => error (cat_lines (map cycle_msg namess)); fun upd_deps name entry G = - foldl (fn (H, parent) => Graph.del_edge (parent, name) H) (G, Graph.imm_preds G name) + Library.foldl (fn (H, parent) => Graph.del_edge (parent, name) H) (G, Graph.imm_preds G name) |> Graph.map_node name (K entry); fun update_node name parents entry G = @@ -132,7 +132,7 @@ fun lookup_thy name = SOME (thy_graph Graph.get_node name) handle Graph.UNDEF _ => NONE; -val known_thy = is_some o lookup_thy; +val known_thy = isSome o lookup_thy; fun check_known_thy name = known_thy name orelse (warning ("Unknown theory " ^ quote name); false); fun if_known_thy f name = if check_known_thy name then f name else (); @@ -146,7 +146,7 @@ (* access deps *) -val lookup_deps = apsome #1 o lookup_thy; +val lookup_deps = Option.map #1 o lookup_thy; val get_deps = #1 o get_thy; fun change_deps name f = change_thy name (fn (deps, x) => (f deps, x)); @@ -158,7 +158,7 @@ NONE => [] | SOME {master, files, ...} => (case master of SOME m => [#1 (ThyLoad.get_thy m)] | NONE => []) @ - (mapfilter (apsome #1 o #2) files)); + (List.mapPartial (Option.map #1 o #2) files)); fun get_preds name = (thy_graph Graph.imm_preds name) handle Graph.UNDEF _ => @@ -197,7 +197,7 @@ fun outdate_thy name = if is_finished name orelse is_outdated name then () - else (change_deps name (apsome (fn {present, outdated = _, master, files} => + else (change_deps name (Option.map (fn {present, outdated = _, master, files} => make_deps present true master files)); perform Outdate name); fun check_unfinished name = @@ -207,12 +207,12 @@ in fun touch_thys names = - seq outdate_thy (thy_graph Graph.all_succs (mapfilter check_unfinished names)); + List.app outdate_thy (thy_graph Graph.all_succs (List.mapPartial check_unfinished names)); fun touch_thy name = touch_thys [name]; fun touch_child_thys name = touch_thys (thy_graph Graph.imm_succs name); -fun touch_all_thys () = seq outdate_thy (get_names ()); +fun touch_all_thys () = List.app outdate_thy (get_names ()); end; @@ -227,9 +227,9 @@ (* load_file *) -val opt_path = apsome (Path.dir o fst o ThyLoad.get_thy); -fun opt_path' (d : deps option) = if_none (apsome (opt_path o #master) d) NONE; -fun opt_path'' d = if_none (apsome opt_path' d) NONE; +val opt_path = Option.map (Path.dir o fst o ThyLoad.get_thy); +fun opt_path' (d : deps option) = getOpt (Option.map (opt_path o #master) d, NONE); +fun opt_path'' d = getOpt (Option.map opt_path' d, NONE); local @@ -241,7 +241,7 @@ | provide _ _ _ NONE = NONE; fun run_file path = - (case apsome PureThy.get_name (Context.get_context ()) of + (case Option.map PureThy.get_name (Context.get_context ()) of NONE => (ThyLoad.load_file NONE path; ()) | SOME name => (case lookup_deps name of SOME deps => change_deps name (provide path name @@ -283,7 +283,7 @@ else #1 (ThyLoad.deps_thy dir name ml); val files = get_files name; - val missing_files = mapfilter (fn (path, NONE) => SOME (Path.pack path) | _ => NONE) files; + val missing_files = List.mapPartial (fn (path, NONE) => SOME (Path.pack path) | _ => NONE) files; in if null missing_files then () else warning (loader_msg "unresolved dependencies of theory" [name] ^ @@ -344,7 +344,7 @@ val (current, (new_deps, parents)) = current_deps ml strict dir name handle ERROR => error (loader_msg "the error(s) above occurred while examining theory" [name] ^ (if null initiators then "" else required_by "\n" initiators)); - val dir' = if_none (opt_path'' new_deps) dir; + val dir' = getOpt (opt_path'' new_deps, dir); val (visited', parent_results) = foldl_map (req_parent dir') (name :: visited, parents); val thy = if not really then SOME (get_theory name) else NONE; @@ -390,7 +390,7 @@ else let val succs = thy_graph Graph.all_succs [name] in priority (loader_msg "removing" succs); - seq (perform Remove) succs; + List.app (perform Remove) succs; change_thys (Graph.del_nodes succs) end; @@ -401,20 +401,20 @@ let val bparents = map base_of_path parents; val dir' = opt_path'' (lookup_deps name); - val dir = if_none dir' Path.current; + val dir = getOpt (dir',Path.current); val assert_thy = if upd then quiet_update_thy' true dir else weak_use_thy dir; val _ = check_unfinished error name; - val _ = seq assert_thy parents; + val _ = List.app assert_thy parents; val theory = PureThy.begin_theory name (map get_theory bparents); val deps = if known_thy name then get_deps name else (init_deps NONE (map #1 paths)); (*records additional ML files only!*) - val use_paths = mapfilter (fn (x, true) => SOME x | _ => NONE) paths; + val use_paths = List.mapPartial (fn (x, true) => SOME x | _ => NONE) paths; val _ = change_thys (update_node name bparents (deps, SOME (Theory.copy theory))); val theory' = theory |> present dir' name bparents paths; val _ = put_theory name (Theory.copy theory'); - val ((), theory'') = Context.pass_theory theory' (seq (load_file false)) use_paths; + val ((), theory'') = Context.pass_theory theory' (List.app (load_file false)) use_paths; val _ = put_theory name (Theory.copy theory''); in theory'' end; @@ -445,7 +445,7 @@ fun get_variant (x, y_name) = if Theory.eq_thy (x, get_theory y_name) then NONE else SOME y_name; - val variants = mapfilter get_variant (parents ~~ parent_names); + val variants = List.mapPartial get_variant (parents ~~ parent_names); fun register G = (Graph.new_node (name, (NONE, SOME theory)) G diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Thy/thy_load.ML --- a/src/Pure/Thy/thy_load.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Thy/thy_load.ML Thu Mar 03 12:43:01 2005 +0100 @@ -75,7 +75,7 @@ fun find_ext rel_path ext = let val ext_path = Path.expand (Path.ext ext rel_path) - in apsome (fn info => (File.full_path ext_path, info)) (File.info ext_path) end; + in Option.map (fn info => (File.full_path ext_path, info)) (File.info ext_path) end; fun find_dir dir = get_first (find_ext (Path.append dir path)) ml_exts; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Thy/thy_parse.ML --- a/src/Pure/Thy/thy_parse.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Thy/thy_parse.ML Thu Mar 03 12:43:01 2005 +0100 @@ -175,7 +175,7 @@ fun mk_triple1 ((x, y), z) = mk_triple (x, y, z); fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z); -fun split_decls l = flat (map (fn (xs, y) => map (rpair y) xs) l); +fun split_decls l = List.concat (map (fn (xs, y) => map (rpair y) xs) l); (* names *) @@ -299,7 +299,7 @@ optional ("=" $$-- typ >> SOME) NONE -- opt_infix >> mk_type_decl; val type_decls = - repeat1 (old_type_decl || type_decl) >> (mk_type_decls o flat); + repeat1 (old_type_decl || type_decl) >> (mk_type_decls o List.concat); (* consts *) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/Thy/thy_scan.ML --- a/src/Pure/Thy/thy_scan.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/Thy/thy_scan.ML Thu Mar 03 12:43:01 2005 +0100 @@ -45,7 +45,7 @@ (* line numbering *) -val incr = apsome (fn n:int => n + 1); +val incr = Option.map (fn n:int => n + 1); fun incr_line scan = Scan.depend (fn n => scan >> pair (incr n)); val keep_line = Scan.lift; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/axclass.ML --- a/src/Pure/axclass.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/axclass.ML Thu Mar 03 12:43:01 2005 +0100 @@ -64,7 +64,7 @@ val is_def = Logic.is_equals o #prop o rep_thm; fun witnesses thy names thms = - PureThy.get_thmss thy (map (rpair NONE) names) @ thms @ filter is_def (map snd (axioms_of thy)); + PureThy.get_thmss thy (map (rpair NONE) names) @ thms @ List.filter is_def (map snd (axioms_of thy)); @@ -248,7 +248,7 @@ [] => [] | [(_, S)] => if Sign.subsort class_sign ([class], S) then S else err_bad_axsort name class | _ => err_bad_tfrees name); - val axS = Sign.certify_sort class_sign (flat (map axm_sort axms)); + val axS = Sign.certify_sort class_sign (List.concat (map axm_sort axms)); val int_axm = Logic.close_form o map_term_tfrees (K (aT axS)); fun inclass c = Logic.mk_inclass (aT axS, c); @@ -288,7 +288,7 @@ fun class_axms sign = let val classes = Sign.classes sign in map (Thm.class_triv sign) classes @ - mapfilter (apsome #intro o lookup_axclass_info_sg sign) classes + List.mapPartial (Option.map #intro o lookup_axclass_info_sg sign) classes end; @@ -326,7 +326,7 @@ fun axclass_tac thms = let - val defs = filter is_def thms; + val defs = List.filter is_def thms; val non_defs = filter_out is_def thms; in intro_classes_tac [] THEN @@ -338,7 +338,7 @@ (* old-style provers *) fun prove mk_prop str_of sign prop thms usr_tac = - (Tactic.prove sign [] [] (mk_prop prop) (K (axclass_tac thms THEN (if_none usr_tac all_tac))) + (Tactic.prove sign [] [] (mk_prop prop) (K (axclass_tac thms THEN (getOpt (usr_tac,all_tac)))) handle ERROR => error ("The error(s) above occurred while trying to prove " ^ quote (str_of sign prop))) |> Drule.standard; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/codegen.ML --- a/src/Pure/codegen.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/codegen.ML Thu Mar 03 12:43:01 2005 +0100 @@ -90,7 +90,7 @@ | args_of (Ignore :: ms) (_ :: xs) = args_of ms xs | args_of (_ :: ms) xs = args_of ms xs; -fun num_args x = length (filter is_arg x); +fun num_args x = length (List.filter is_arg x); (**** theory data ****) @@ -225,7 +225,7 @@ fun mk_parser (a, p) = (if a = "" then Scan.succeed "" else Args.$$$ a) |-- p; val code_attr = - Attrib.syntax (Scan.depend (fn thy => foldr op || (map mk_parser + Attrib.syntax (Scan.depend (fn thy => Library.foldr op || (map mk_parser (#attrs (CodegenData.get thy)), Scan.fail) >> pair thy)); @@ -242,7 +242,7 @@ fun preprocess thy ths = let val {preprocs, ...} = CodegenData.get thy - in foldl (fn (ths, (_, f)) => f thy ths) (ths, preprocs) end; + in Library.foldl (fn (ths, (_, f)) => f thy ths) (ths, preprocs) end; fun unfold_attr (thy, eqn) = let @@ -257,7 +257,7 @@ (**** associate constants with target language code ****) -fun gen_assoc_consts prep_type xs thy = foldl (fn (thy, (s, tyopt, syn)) => +fun gen_assoc_consts prep_type xs thy = Library.foldl (fn (thy, (s, tyopt, syn)) => let val sg = sign_of thy; val {codegens, tycodegens, consts, types, attrs, preprocs, test_params} = @@ -289,7 +289,7 @@ (**** associate types with target language types ****) -fun assoc_types xs thy = foldl (fn (thy, (s, syn)) => +fun assoc_types xs thy = Library.foldl (fn (thy, (s, syn)) => let val {codegens, tycodegens, consts, types, attrs, preprocs, test_params} = CodegenData.get thy; @@ -329,7 +329,7 @@ | (ctrl, s') => (if ctrl then "ctrl_" ^ s' else s') :: check_str zs) | (ys, zs) => implode ys :: check_str zs); val s' = space_implode "_" - (flat (map (check_str o Symbol.explode) (NameSpace.unpack s))) + (List.concat (map (check_str o Symbol.explode) (NameSpace.unpack s))) in if Symbol.is_ascii_letter (hd (explode s')) then s' else "id_" ^ s' end; @@ -344,15 +344,15 @@ fun rename_terms ts = let - val names = foldr add_term_names + val names = Library.foldr add_term_names (ts, map (fst o fst) (Drule.vars_of_terms ts)); val reserved = names inter ThmDatabase.ml_reserved; - val (illegal, alt_names) = split_list (mapfilter (fn s => + val (illegal, alt_names) = split_list (List.mapPartial (fn s => let val s' = mk_id s in if s = s' then NONE else SOME (s, s') end) names) val ps = (reserved @ illegal) ~~ variantlist (map (suffix "'") reserved @ alt_names, names); - fun rename_id s = if_none (assoc (ps, s)) s; + fun rename_id s = getOpt (assoc (ps, s), s); fun rename (Var ((a, i), T)) = Var ((rename_id a, i), T) | rename (Free (a, T)) = Free (rename_id a, T) @@ -371,12 +371,12 @@ fun is_instance thy T1 T2 = Sign.typ_instance (sign_of thy) (T1, Type.varifyT T2); -fun get_assoc_code thy s T = apsome snd (find_first (fn ((s', T'), _) => +fun get_assoc_code thy s T = Option.map snd (find_first (fn ((s', T'), _) => s = s' andalso is_instance thy T T') (#consts (CodegenData.get thy))); fun get_defn thy s T = let - val axms = flat (map (Symtab.dest o #axioms o Theory.rep_theory) + val axms = List.concat (map (Symtab.dest o #axioms o Theory.rep_theory) (thy :: Theory.ancestors_of thy)); fun prep_def def = (case preprocess thy [def] of [def'] => prop_of def' | _ => error "get_defn: bad preprocessor"); @@ -387,11 +387,11 @@ val (s', T') = dest_Const c in if s = s' then SOME (T', (args, rhs)) else NONE end handle TERM _ => NONE; - val defs = mapfilter (fn (name, t) => apsome (pair name) (dest t)) axms; + val defs = List.mapPartial (fn (name, t) => Option.map (pair name) (dest t)) axms; val i = find_index (is_instance thy T o fst o snd) defs in if i >= 0 then - let val (name, (T', (args, _))) = nth_elem (i, defs) + let val (name, (T', (args, _))) = List.nth (defs, i) in case dest (prep_def (Thm.get_axiom thy name)) of NONE => NONE | SOME (T'', p as (args', rhs)) => @@ -443,8 +443,8 @@ val (Ts, _) = strip_type (fastype_of t); val j = i - length ts in - foldr (fn (T, t) => Abs ("x", T, t)) - (take (j, Ts), list_comb (t, ts @ map Bound (j-1 downto 0))) + Library.foldr (fn (T, t) => Abs ("x", T, t)) + (Library.take (j, Ts), list_comb (t, ts @ map Bound (j-1 downto 0))) end; fun mk_app _ p [] = p @@ -583,7 +583,7 @@ val strip_tname = implode o tl o explode; fun pretty_list xs = Pretty.block (Pretty.str "[" :: - flat (separate [Pretty.str ",", Pretty.brk 1] (map single xs)) @ + List.concat (separate [Pretty.str ",", Pretty.brk 1] (map single xs)) @ [Pretty.str "]"]); fun mk_type p (TVar ((s, i), _)) = Pretty.str @@ -598,7 +598,7 @@ | mk_term_of sg p (TFree (s, _)) = Pretty.str (strip_tname s ^ "F") | mk_term_of sg p (Type (s, Ts)) = (if p then parens else I) (Pretty.block (separate (Pretty.brk 1) (Pretty.str ("term_of_" ^ mk_type_id sg s) :: - flat (map (fn T => [mk_term_of sg true T, mk_type true T]) Ts)))); + List.concat (map (fn T => [mk_term_of sg true T, mk_type true T]) Ts)))); (**** Test data generators ****) @@ -639,7 +639,7 @@ Pretty.brk 1, Pretty.str "then NONE", Pretty.brk 1, Pretty.str "else ", Pretty.block [Pretty.str "SOME ", Pretty.block (Pretty.str "[" :: - flat (separate [Pretty.str ",", Pretty.brk 1] + List.concat (separate [Pretty.str ",", Pretty.brk 1] (map (fn (s, T) => [Pretty.block [Pretty.str ("(" ^ Library.quote (Symbol.escape s) ^ ","), Pretty.brk 1, mk_app false (mk_term_of sg false T) @@ -666,8 +666,8 @@ val (gi, frees) = Logic.goal_params (prop_of (snd (snd (Proof.get_goal (Toplevel.proof_of st))))) i; val gi' = ObjectLogic.atomize_term sg (map_term_types - (map_type_tfree (fn p as (s, _) => if_none (assoc (tvinsts, s)) - (if_none default_type (TFree p)))) (subst_bounds (frees, strip gi))); + (map_type_tfree (fn p as (s, _) => getOpt (assoc (tvinsts, s), + getOpt (default_type,TFree p)))) (subst_bounds (frees, strip gi))); in case test_term (Toplevel.theory_of st) size iterations gi' of NONE => writeln "No counterexamples found." | SOME cex => writeln ("Counterexample found:\n" ^ @@ -736,7 +736,7 @@ ("default_type", P.typ >> set_default_type)]; val parse_test_params = P.short_ident :-- (fn s => - P.$$$ "=" |-- if_none (assoc (params, s)) Scan.fail) >> snd; + P.$$$ "=" |-- getOpt (assoc (params, s), Scan.fail)) >> snd; fun parse_tyinst xs = (P.type_ident --| P.$$$ "=" -- P.typ >> (fn (v, s) => fn sg => @@ -757,8 +757,8 @@ ( parse_test_params >> (fn f => fn sg => apfst (f sg)) || parse_tyinst) --| P.$$$ "]") -- Scan.optional P.nat 1 >> (fn (ps, g) => Toplevel.keep (fn st => - test_goal (app (if_none (apsome - (map (fn f => f (Toplevel.sign_of st))) ps) []) + test_goal (app (getOpt (Option.map + (map (fn f => f (Toplevel.sign_of st))) ps, [])) (get_test_params (Toplevel.theory_of st), [])) g st))); val parsers = [assoc_typeP, assoc_constP, generate_codeP, test_paramsP, testP]; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/display.ML --- a/src/Pure/display.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/display.ML Thu Mar 03 12:43:01 2005 +0100 @@ -122,7 +122,7 @@ (Seq.print (fn _ => print_thm) 100000 thseq; thseq); (*Print and return a list of theorems, separated by blank lines. *) -fun prths ths = (seq (fn th => (print_thm th; writeln "")) ths; ths); +fun prths ths = (List.app (fn th => (print_thm th; writeln "")) ths; ths); (* other printing commands *) @@ -210,7 +210,7 @@ if not syn then NONE else SOME (prt_typ (Type (t, []))); - val pretty_arities = flat o map (fn (t, ars) => map (prt_arity t) ars); + val pretty_arities = List.concat o map (fn (t, ars) => map (prt_arity t) ars); fun pretty_const (c, (ty, _)) = Pretty.block [Pretty.str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty]; @@ -238,13 +238,13 @@ in [Pretty.strs ("stamps:" :: Sign.stamp_names_of sg), Pretty.strs ("data:" :: Sign.data_kinds data), - Pretty.strs ["name prefix:", NameSpace.pack (if_none path ["-"])], + Pretty.strs ["name prefix:", NameSpace.pack (getOpt (path,["-"]))], Pretty.big_list "name spaces:" (map pretty_name_space spcs), Pretty.big_list "classes:" (map pretty_classrel (Graph.dest classes)), pretty_default default, pretty_witness witness, - Pretty.big_list "syntactic types:" (mapfilter (pretty_type true) tdecls), - Pretty.big_list "logical types:" (mapfilter (pretty_type false) tdecls), + Pretty.big_list "syntactic types:" (List.mapPartial (pretty_type true) tdecls), + Pretty.big_list "logical types:" (List.mapPartial (pretty_type false) tdecls), Pretty.big_list "type arities:" (pretty_arities (Symtab.dest arities)), Pretty.big_list "consts:" (map pretty_const cnsts), Pretty.big_list "finals consts:" (map pretty_final finals), @@ -282,7 +282,7 @@ | add_vars (t $ u, env) = add_vars (u, add_vars (t, env)) | add_vars (_, env) = env; -fun add_varsT (Type (_, Ts), env) = foldr add_varsT (Ts, env) +fun add_varsT (Type (_, Ts), env) = Library.foldr add_varsT (Ts, env) | add_varsT (TFree (x, S), env) = ins_entry (S, (x, ~1)) env | add_varsT (TVar (xi, S), env) = ins_entry (S, xi) env; @@ -334,7 +334,7 @@ (if main then [Pretty.term pp B] else []) @ (if ngoals = 0 then [Pretty.str "No subgoals!"] else if ngoals > maxgoals then - pretty_subgoals (take (maxgoals, As)) @ + pretty_subgoals (Library.take (maxgoals, As)) @ (if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")] else []) else pretty_subgoals As) @ diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/drule.ML --- a/src/Pure/drule.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/drule.ML Thu Mar 03 12:43:01 2005 +0100 @@ -221,7 +221,7 @@ let fun is_tv ((a, _), _) = (case Symbol.explode a of "'" :: _ => true | _ => false); - val (tvs, vs) = partition is_tv insts; + val (tvs, vs) = List.partition is_tv insts; fun readT (ixn, st) = let val S = case rsorts ixn of SOME S => S | NONE => absent ixn; val T = Sign.read_typ (sign,sorts) st; @@ -340,16 +340,16 @@ val {sign, prop, maxidx, ...} = Thm.rep_thm thm; fun elim (th, (x, T)) = Thm.forall_elim (Thm.cterm_of sign (Var ((x, maxidx + 1), T))) th; val vs = Term.strip_all_vars prop; - in foldl elim (thm, Term.variantlist (map #1 vs, []) ~~ map #2 vs) end; + in Library.foldl elim (thm, Term.variantlist (map #1 vs, []) ~~ map #2 vs) end; (*Specialization over a list of cterms*) -fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th); +fun forall_elim_list cts th = Library.foldr (uncurry forall_elim) (rev cts, th); (* maps A1,...,An |- B to [| A1;...;An |] ==> B *) -fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th); +fun implies_intr_list cAs th = Library.foldr (uncurry implies_intr) (cAs,th); (* maps [| A1;...;An |] ==> B and [A1,...,An] to B *) -fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths); +fun implies_elim_list impth ths = Library.foldl (uncurry implies_elim) (impth,ths); (* maps |- B to A1,...,An |- B *) fun impose_hyps chyps th = @@ -364,13 +364,13 @@ fun zero_var_indexes th = let val {prop,sign,tpairs,...} = rep_thm th; val (tpair_l, tpair_r) = Library.split_list tpairs; - val vars = foldr add_term_vars - (tpair_r, foldr add_term_vars (tpair_l, (term_vars prop))); - val bs = foldl add_new_id ([], map (fn Var((a,_),_)=>a) vars) + val vars = Library.foldr add_term_vars + (tpair_r, Library.foldr add_term_vars (tpair_l, (term_vars prop))); + val bs = Library.foldl add_new_id ([], map (fn Var((a,_),_)=>a) vars) val inrs = - foldr add_term_tvars - (tpair_r, foldr add_term_tvars (tpair_l, (term_tvars prop))); - val nms' = rev(foldl add_new_id ([], map (#1 o #1) inrs)); + Library.foldr add_term_tvars + (tpair_r, Library.foldr add_term_tvars (tpair_l, (term_tvars prop))); + val nms' = rev(Library.foldl add_new_id ([], map (#1 o #1) inrs)); val tye = ListPair.map (fn ((v,rs),a) => (v, TVar((a,0),rs))) (inrs, nms') val ctye = map (fn (v,T) => (v,ctyp_of sign T)) tye; @@ -423,16 +423,16 @@ let val fth = freezeT th val {prop, tpairs, sign, ...} = rep_thm fth in - case foldr add_term_vars (prop :: Thm.terms_of_tpairs tpairs, []) of + case Library.foldr add_term_vars (prop :: Thm.terms_of_tpairs tpairs, []) of [] => (fth, fn i => fn x => x) (*No vars: nothing to do!*) | vars => let fun newName (Var(ix,_), pairs) = let val v = gensym (string_of_indexname ix) in ((ix,v)::pairs) end; - val alist = foldr newName (vars,[]) + val alist = Library.foldr newName (vars,[]) fun mk_inst (Var(v,T)) = (cterm_of sign (Var(v,T)), - cterm_of sign (Free(the (assoc(alist,v)), T))) + cterm_of sign (Free(valOf (assoc(alist,v)), T))) val insts = map mk_inst vars fun thaw i th' = (*i is non-negative increment for Var indexes*) th' |> forall_intr_list (map #2 insts) @@ -446,17 +446,17 @@ let val fth = freezeT th val {prop, tpairs, sign, ...} = rep_thm fth in - case foldr add_term_vars (prop :: Thm.terms_of_tpairs tpairs, []) of + case Library.foldr add_term_vars (prop :: Thm.terms_of_tpairs tpairs, []) of [] => (fth, fn x => x) | vars => let fun newName (Var(ix,_), (pairs,used)) = let val v = variant used (string_of_indexname ix) in ((ix,v)::pairs, v::used) end; - val (alist, _) = foldr newName (vars, ([], foldr add_term_names + val (alist, _) = Library.foldr newName (vars, ([], Library.foldr add_term_names (prop :: Thm.terms_of_tpairs tpairs, []))) fun mk_inst (Var(v,T)) = (cterm_of sign (Var(v,T)), - cterm_of sign (Free(the (assoc(alist,v)), T))) + cterm_of sign (Free(valOf (assoc(alist,v)), T))) val insts = map mk_inst vars fun thaw th' = th' |> forall_intr_list (map #2 insts) @@ -641,7 +641,7 @@ fun abs_def thm = let val (_, cvs) = strip_comb (fst (dest_equals (cprop_of thm))); - val thm' = foldr (fn (ct, thm) => Thm.abstract_rule + val thm' = Library.foldr (fn (ct, thm) => Thm.abstract_rule (case term_of ct of Var ((a, _), _) => a | Free (a, _) => a | _ => "x") ct thm) (cvs, thm) in transitive @@ -835,7 +835,7 @@ in (sign', tye', maxi') end; in fun cterm_instantiate ctpairs0 th = - let val (sign,tye,_) = foldr add_types (ctpairs0, (Thm.sign_of_thm th, Vartab.empty, 0)) + let val (sign,tye,_) = Library.foldr add_types (ctpairs0, (Thm.sign_of_thm th, Vartab.empty, 0)) fun instT(ct,cu) = let val inst = cterm_of sign o subst_TVars_Vartab tye o term_of in (inst ct, inst cu) end @@ -862,7 +862,7 @@ handle THM _ => raise THM("equal_abs_elim", 0, [eqth]); (*Calling equal_abs_elim with multiple terms*) -fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th); +fun equal_abs_elim_list cts th = Library.foldr (uncurry equal_abs_elim) (rev cts, th); (*** Goal (PROP A) <==> PROP A ***) @@ -896,8 +896,8 @@ (* collect vars in left-to-right order *) -fun tvars_of_terms ts = rev (foldl Term.add_tvars ([], ts)); -fun vars_of_terms ts = rev (foldl Term.add_vars ([], ts)); +fun tvars_of_terms ts = rev (Library.foldl Term.add_tvars ([], ts)); +fun vars_of_terms ts = rev (Library.foldl Term.add_vars ([], ts)); fun tvars_of thm = tvars_of_terms [prop_of thm]; fun vars_of thm = vars_of_terms [prop_of thm]; @@ -909,8 +909,8 @@ let fun err msg = raise TYPE ("instantiate': " ^ msg, - mapfilter (apsome Thm.typ_of) cTs, - mapfilter (apsome Thm.term_of) cts); + List.mapPartial (Option.map Thm.typ_of) cTs, + List.mapPartial (Option.map Thm.term_of) cts); fun inst_of (v, ct) = (Thm.cterm_of (#sign (Thm.rep_cterm ct)) (Var v), ct) @@ -941,7 +941,7 @@ | rename_bvars vs thm = let val {sign, prop, ...} = rep_thm thm; - fun ren (Abs (x, T, t)) = Abs (if_none (assoc (vs, x)) x, T, ren t) + fun ren (Abs (x, T, t)) = Abs (getOpt (assoc (vs, x), x), T, ren t) | ren (t $ u) = ren t $ ren u | ren t = t; in equal_elim (reflexive (cterm_of sign (ren prop))) thm end; @@ -955,7 +955,7 @@ fun rename [] t = ([], t) | rename (x' :: xs) (Abs (x, T, t)) = let val (xs', t') = rename xs t - in (xs', Abs (if_none x' x, T, t')) end + in (xs', Abs (getOpt (x',x), T, t')) end | rename xs (t $ u) = let val (xs', t') = rename xs t; @@ -991,7 +991,7 @@ fun tfrees_of thm = let val {hyps, prop, ...} = Thm.rep_thm thm - in foldr Term.add_term_tfree_names (prop :: hyps, []) end; + in Library.foldr Term.add_term_tfree_names (prop :: hyps, []) end; fun tvars_intr_list tfrees thm = Thm.varifyT' (tfrees_of thm \\ tfrees) thm; @@ -1002,7 +1002,7 @@ fun incr_indexes_wrt is cTs cts thms = let val maxidx = - foldl Int.max (~1, is @ + Library.foldl Int.max (~1, is @ map (maxidx_of_typ o #T o Thm.rep_ctyp) cTs @ map (#maxidx o Thm.rep_cterm) cts @ map (#maxidx o Thm.rep_thm) thms); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/envir.ML --- a/src/Pure/envir.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/envir.ML Thu Mar 03 12:43:01 2005 +0100 @@ -197,8 +197,8 @@ | fast Ts (Const (_, T)) = T | fast Ts (Free (_, T)) = T | fast Ts (Bound i) = - (nth_elem (i, Ts) - handle LIST _=> raise TERM ("fastype: Bound", [Bound i])) + (List.nth (Ts, i) + handle Subscript => raise TERM ("fastype: Bound", [Bound i])) | fast Ts (Var (_, T)) = T | fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u in fast end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/fact_index.ML --- a/src/Pure/fact_index.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/fact_index.ML Thu Mar 03 12:43:01 2005 +0100 @@ -32,7 +32,7 @@ fun index_thm pred (idx, th) = let val {hyps, prop, ...} = Thm.rep_thm th - in foldl (index_term pred) (index_term pred (idx, prop), hyps) end; + in Library.foldl (index_term pred) (index_term pred (idx, prop), hyps) end; (* build index *) @@ -47,10 +47,10 @@ fun add pred (Index {next, facts, consts, frees}, (name, ths)) = let fun upd (tab, a) = Symtab.update_multi ((a, (next, (name, ths))), tab); - val (cs, xs) = foldl (index_thm pred) (([], []), ths); + val (cs, xs) = Library.foldl (index_thm pred) (([], []), ths); in Index {next = next + 1, facts = (name, ths) :: facts, - consts = foldl upd (consts, cs), frees = foldl upd (frees, xs)} + consts = Library.foldl upd (consts, cs), frees = Library.foldl upd (frees, xs)} end; @@ -64,7 +64,7 @@ else intersect (xxs, ys); fun intersects [xs] = xs - | intersects xss = if exists null xss then [] else foldl intersect (hd xss, tl xss); + | intersects xss = if exists null xss then [] else Library.foldl intersect (hd xss, tl xss); fun find ([], []) (Index {facts, ...}) = facts | find (cs, xs) (Index {consts, frees, ...}) = diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/goals.ML --- a/src/Pure/goals.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/goals.ML Thu Mar 03 12:43:01 2005 +0100 @@ -236,7 +236,7 @@ let val {space, locales, ...} = LocalesData.get thy; val name = NameSpace.intern space xname; - in apsome (pair name) (get_locale thy name) end; + in Option.map (pair name) (get_locale thy name) end; (* access scope *) @@ -307,8 +307,8 @@ let val cur_sc = get_scope_sg sg; val rule_lists = map (#rules o snd) cur_sc; val def_lists = map (#defs o snd) cur_sc; - val rules = map snd (foldr (op union) (rule_lists, [])); - val defs = map snd (foldr (op union) (def_lists, [])); + val rules = map snd (Library.foldr (op union) (rule_lists, [])); + val defs = map snd (Library.foldr (op union) (def_lists, [])); val defnrules = rules @ defs; in hyps subset defnrules @@ -394,9 +394,9 @@ fun read_cterm sign = let val cur_sc = get_scope_sg sign; val defaults = map (#defaults o snd) cur_sc; - val envS = flat (map #1 defaults); - val envT = flat (map #2 defaults); - val used = flat (map #3 defaults); + val envS = List.concat (map #1 defaults); + val envT = List.concat (map #2 defaults); + val used = List.concat (map #3 defaults); fun def_sort (x, ~1) = assoc (envS, x) | def_sort _ = NONE; fun def_type (x, ~1) = assoc (envT, x) @@ -411,7 +411,7 @@ fun collect_consts sg = let val cur_sc = get_scope_sg sg; val locale_list = map snd cur_sc; - val const_list = flat (map #consts locale_list) + val const_list = List.concat (map #consts locale_list) in map fst const_list end; (* filter out the Free's in a term *) @@ -447,8 +447,8 @@ fun abs_o (t, (x as Free(v,T))) = all(T) $ Abs(v, T, abstract_over (x,t)) | abs_o (_ , _) = error ("Can't be: abs_over_free"); val diffl' = map (fn (Free (s, T)) => s) diffl; - val defaults' = (#1 defaults, filter (fn x => not((fst x) mem diffl')) (#2 defaults), #3 defaults) - in (defaults', (s, foldl abs_o (term, diffl))) end; + val defaults' = (#1 defaults, List.filter (fn x => not((fst x) mem diffl')) (#2 defaults), #3 defaults) + in (defaults', (s, Library.foldl abs_o (term, diffl))) end; (* assume a definition, i.e assume the cterm of a definiton term and then eliminate the binding !!, so that the def can be applied as rewrite. The meta hyp will still contain !! *) @@ -499,8 +499,8 @@ val loc_consts = map #1 loc_consts_syn; val loc_consts = old_loc_consts @ loc_consts; val loc_syn = map #2 loc_consts_syn; - val nosyn = old_nosyn @ (map (#1 o #1) (filter (fn x => (#3(#2 x)) = NoSyn) loc_consts_syn)); - val loc_trfuns = mapfilter #3 loc_consts_syn; + val nosyn = old_nosyn @ (map (#1 o #1) (List.filter (fn x => (#3(#2 x)) = NoSyn) loc_consts_syn)); + val loc_trfuns = List.mapPartial #3 loc_consts_syn; (* 1st stage: syntax_thy *) @@ -574,7 +574,7 @@ end | compare_var_sides (_,_) = locale_error ("Definitions must use the == relation") - val check = foldl compare_var_sides (true, loc_defs) + val check = Library.foldl compare_var_sides (true, loc_defs) in if check then [] else ["nonfixed variable on right hand side of a locale definition in locale " ^ quote name] end; @@ -620,7 +620,7 @@ fun pretty_term sign = if (is_open_loc_sg sign) then let val locale_list = map snd(get_scope_sg sign); - val nosyn = flat (map #nosyn locale_list); + val nosyn = List.concat (map #nosyn locale_list); val str_list = (collect_consts sign) \\ nosyn in Sign.pretty_term sign o (const_ssubst_list str_list) end @@ -715,7 +715,7 @@ val sign = Theory.sign_of thy; val used_As = A_set inter #hyps(rep_thm(th)); val ctl = map (cterm_of sign) used_As - in foldl (fn (x, y) => Thm.implies_intr y x) (th, ctl) + in Library.foldl (fn (x, y) => Thm.implies_intr y x) (th, ctl) end; fun standard_some thy A_set th = @@ -776,7 +776,7 @@ ("Additional hypotheses:\n" ^ cat_lines (map (Sign.string_of_term sign) - (filter (fn x => (not (in_locale [x] sign))) + (List.filter (fn x => (not (in_locale [x] sign))) hyps))) else if Pattern.matches (Sign.tsig_of sign) (Envir.beta_norm (term_of chorn), Envir.beta_norm prop) @@ -796,17 +796,17 @@ case e of THM (msg,i,thms) => (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); - seq print_thm thms) + List.app print_thm thms) | THEORY (msg,thys) => (writeln ("Exception THEORY raised:\n" ^ msg); - seq (Pretty.writeln o Display.pretty_theory) thys) + List.app (Pretty.writeln o Display.pretty_theory) thys) | TERM (msg,ts) => (writeln ("Exception TERM raised:\n" ^ msg); - seq (writeln o Sign.string_of_term sign) ts) + List.app (writeln o Sign.string_of_term sign) ts) | TYPE (msg,Ts,ts) => (writeln ("Exception TYPE raised:\n" ^ msg); - seq (writeln o Sign.string_of_typ sign) Ts; - seq (writeln o Sign.string_of_term sign) ts) + List.app (writeln o Sign.string_of_typ sign) Ts; + List.app (writeln o Sign.string_of_term sign) ts) | e => raise e; (*Prints an exception, then fails*) @@ -1146,7 +1146,7 @@ val thy = theory_of_goal (); val consts = map (Sign.intern_const (Theory.sign_of thy)) raw_consts; in - (case filter (is_none o Sign.const_type (Theory.sign_of thy)) consts of + (case List.filter (is_none o Sign.const_type (Theory.sign_of thy)) consts of [] => () | cs => error ("thms_containing: undeclared consts " ^ commas_quote cs)); PureThy.thms_containing_consts thy consts @@ -1159,12 +1159,12 @@ val prems = Logic.prems_of_goal (prop_of (topthm())) i val thy = theory_of_goal(); val thmss = map (PureThy.find_elims thy) prems - in foldl (gen_union eq_nth) ([],thmss) end; + in Library.foldl (gen_union eq_nth) ([],thmss) end; fun findE i j = let val prems = Logic.prems_of_goal (prop_of (topthm())) i val thy = theory_of_goal(); - in PureThy.find_elims thy (nth_elem(j-1, prems)) end; + in PureThy.find_elims thy (List.nth(prems, j-1)) end; end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/library.ML --- a/src/Pure/library.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/library.ML Thu Mar 03 12:43:01 2005 +0100 @@ -36,11 +36,7 @@ val stamp: unit -> stamp (*options*) - val the: 'a option -> 'a - val if_none: 'a option -> 'a -> 'a - val is_some: 'a option -> bool val is_none: 'a option -> bool - val apsome: ('a -> 'b) -> 'a option -> 'b option exception ERROR val try: ('a -> 'b) -> 'a -> 'b option val can: ('a -> 'b) -> 'a -> bool @@ -77,45 +73,31 @@ val conditional: bool -> (unit -> unit) -> unit (*lists*) - exception LIST of string - val hd: 'a list -> 'a - val tl: 'a list -> 'a list + exception UnequalLengths val cons: 'a -> 'a list -> 'a list val single: 'a -> 'a list val append: 'a list -> 'a list -> 'a list - val rev_append: 'a list -> 'a list -> 'a list val apply: ('a -> 'a) list -> 'a -> 'a - val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a - val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b val foldr1: ('a * 'a -> 'a) -> 'a list -> 'a val fold: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b val fold_rev: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b val foldl_map: ('a * 'b -> 'a * 'c) -> 'a * 'b list -> 'a * 'c list val foldln: ('a * int -> 'b -> 'b) -> 'a list -> 'b -> 'b - val length: 'a list -> int - val take: int * 'a list -> 'a list - val drop: int * 'a list -> 'a list + val unflat: 'a list list -> 'b list -> 'b list list val splitAt: int * 'a list -> 'a list * 'a list val dropwhile: ('a -> bool) -> 'a list -> 'a list - val nth_elem: int * 'a list -> 'a val map_nth_elem: int -> ('a -> 'a) -> 'a list -> 'a list - val last_elem: 'a list -> 'a val split_last: 'a list -> 'a list * 'a val nth_update: 'a -> int * 'a list -> 'a list val find_index: ('a -> bool) -> 'a list -> int val find_index_eq: ''a -> ''a list -> int val find_first: ('a -> bool) -> 'a list -> 'a option val get_first: ('a -> 'b option) -> 'a list -> 'b option - val flat: 'a list list -> 'a list - val unflat: 'a list list -> 'b list -> 'b list list - val seq: ('a -> unit) -> 'a list -> unit val separate: 'a -> 'a list -> 'a list val replicate: int -> 'a -> 'a list val multiply: 'a list * 'a list list -> 'a list list val product: 'a list -> 'b list -> ('a * 'b) list - val filter: ('a -> bool) -> 'a list -> 'a list val filter_out: ('a -> bool) -> 'a list -> 'a list - val mapfilter: ('a -> 'b option) -> 'a list -> 'b list val map2: ('a * 'b -> 'c) -> 'a list * 'b list -> 'c list val exists2: ('a * 'b -> bool) -> 'a list * 'b list -> bool val forall2: ('a * 'b -> bool) -> 'a list * 'b list -> bool @@ -262,14 +244,36 @@ (*misc*) val make_keylist: ('a -> 'b) -> 'a list -> ('a * 'b) list val keyfilter: ('a * ''b) list -> ''b -> 'a list - val partition: ('a -> bool) -> 'a list -> 'a list * 'a list val partition_eq: ('a * 'a -> bool) -> 'a list -> 'a list list val partition_list: (int -> 'a -> bool) -> int -> int -> 'a list -> 'a list list val gensym: string -> string val scanwords: (string -> bool) -> string list -> string list end; -structure Library: LIBRARY = +signature LIBRARY_CLOSED = +sig + include LIBRARY + + val the: 'a option -> 'a + val if_none: 'a option -> 'a -> 'a + val is_some: 'a option -> bool + val apsome: ('a -> 'b) -> 'a option -> 'b option + + val rev_append: 'a list -> 'a list -> 'a list + val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a + val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b + val take: int * 'a list -> 'a list + val drop: int * 'a list -> 'a list + val nth_elem: int * 'a list -> 'a + val last_elem: 'a list -> 'a + val flat: 'a list list -> 'a list + val seq: ('a -> unit) -> 'a list -> unit + val partition: ('a -> bool) -> 'a list -> 'a list * 'a list + val filter: ('a -> bool) -> 'a list -> 'a list + val mapfilter: ('a -> 'b option) -> 'a list -> 'b list +end; + +structure Library: LIBRARY_CLOSED = struct @@ -312,18 +316,16 @@ (** options **) -fun the opt = valOf opt; +val the = valOf; (*strict!*) fun if_none opt a = getOpt(opt,a); -fun is_some opt = isSome opt; +val is_some = isSome; -fun is_none (SOME _) = false - | is_none NONE = true; +fun is_none opt = not (isSome opt); -fun apsome f (SOME x) = SOME (f x) - | apsome _ NONE = NONE; +val apsome = Option.map; (* exception handling *) @@ -429,13 +431,7 @@ (** lists **) -exception LIST of string; - -fun hd [] = raise LIST "hd" - | hd (x :: _) = x; - -fun tl [] = raise LIST "tl" - | tl (_ :: xs) = xs; +exception UnequalLengths; fun cons x xs = x :: xs; fun single x = [x]; @@ -488,11 +484,6 @@ (* basic list functions *) -(*length of a list, should unquestionably be a standard function*) -local fun length1 (n, []) = n (*TAIL RECURSIVE*) - | length1 (n, x :: xs) = length1 (n + 1, xs) -in fun length l = length1 (0, l) end; - (*take the first n elements from a list*) fun take (n, []) = [] | take (n, x :: xs) = @@ -513,29 +504,24 @@ (*return nth element of a list, where 0 designates the first element; raise EXCEPTION if list too short*) -fun nth_elem NL = - (case drop NL of - [] => raise LIST "nth_elem" - | x :: _ => x); +fun nth_elem (i,xs) = List.nth(xs,i); fun map_nth_elem 0 f (x :: xs) = f x :: xs | map_nth_elem n f (x :: xs) = x :: map_nth_elem (n - 1) f xs - | map_nth_elem _ _ [] = raise LIST "map_nth_elem"; + | map_nth_elem _ _ [] = raise Subscript; (*last element of a list*) -fun last_elem [] = raise LIST "last_elem" - | last_elem [x] = x - | last_elem (_ :: xs) = last_elem xs; +val last_elem = List.last; (*rear decomposition*) -fun split_last [] = raise LIST "split_last" +fun split_last [] = raise Empty | split_last [x] = ([], x) | split_last (x :: xs) = apfst (cons x) (split_last xs); (*update nth element*) fun nth_update x n_xs = (case splitAt n_xs of - (_,[]) => raise LIST "nth_update" + (_,[]) => raise Subscript | (prfx, _ :: sffx') => prfx @ (x :: sffx')) (*find the position of an element in a list*) @@ -565,14 +551,11 @@ let val (ps,qs) = splitAt(length xs,ys) in ps :: unflat xss qs end | unflat [] [] = [] - | unflat _ _ = raise LIST "unflat"; + | unflat _ _ = raise UnequalLengths; (*like Lisp's MAPC -- seq proc [x1, ..., xn] evaluates (proc x1; ...; proc xn) for side effects*) -fun seq (proc: 'a -> unit) : 'a list -> unit = - let fun seqf [] = () - | seqf (x :: xs) = (proc x; seqf xs) - in seqf end; +val seq = List.app; (*separate s [x1, x2, ..., xn] ===> [x1, s, x2, s, ..., s, xn]*) fun separate s (x :: (xs as _ :: _)) = x :: s :: separate s xs @@ -583,7 +566,7 @@ let fun rep (0, xs) = xs | rep (n, xs) = rep (n - 1, x :: xs) in - if n < 0 then raise LIST "replicate" + if n < 0 then raise Subscript else rep (n, []) end; @@ -606,40 +589,38 @@ fun filter_out f = filter (not o f); -fun mapfilter (f: 'a -> 'b option) ([]: 'a list) = [] : 'b list - | mapfilter f (x :: xs) = - (case f x of - NONE => mapfilter f xs - | SOME y => y :: mapfilter f xs); +val mapfilter = List.mapPartial; (* lists of pairs *) +exception UnequalLengths; + fun map2 _ ([], []) = [] | map2 f (x :: xs, y :: ys) = (f (x, y) :: map2 f (xs, ys)) - | map2 _ _ = raise LIST "map2"; + | map2 _ _ = raise UnequalLengths; fun exists2 _ ([], []) = false | exists2 pred (x :: xs, y :: ys) = pred (x, y) orelse exists2 pred (xs, ys) - | exists2 _ _ = raise LIST "exists2"; + | exists2 _ _ = raise UnequalLengths; fun forall2 _ ([], []) = true | forall2 pred (x :: xs, y :: ys) = pred (x, y) andalso forall2 pred (xs, ys) - | forall2 _ _ = raise LIST "forall2"; + | forall2 _ _ = raise UnequalLengths; fun seq2 _ ([], []) = () | seq2 f (x :: xs, y :: ys) = (f (x, y); seq2 f (xs, ys)) - | seq2 _ _ = raise LIST "seq2"; + | seq2 _ _ = raise UnequalLengths; (*combine two lists forming a list of pairs: [x1, ..., xn] ~~ [y1, ..., yn] ===> [(x1, y1), ..., (xn, yn)]*) fun [] ~~ [] = [] | (x :: xs) ~~ (y :: ys) = (x, y) :: (xs ~~ ys) - | _ ~~ _ = raise LIST "~~"; + | _ ~~ _ = raise UnequalLengths; (*inverse of ~~; the old 'split': [(x1, y1), ..., (xn, yn)] ===> ([x1, ..., xn], [y1, ..., yn])*) -fun split_list (l: ('a * 'b) list) = (map #1 l, map #2 l); +val split_list = ListPair.unzip; fun equal_lists eq (xs, ys) = length xs = length ys andalso forall2 eq (xs, ys); @@ -751,7 +732,7 @@ fun nth_elem_string (i, str) = (case try String.substring (str, i, 1) of SOME s => s - | NONE => raise LIST "nth_elem_string"); + | NONE => raise Subscript); fun foldl_string f (x0, str) = let @@ -822,7 +803,7 @@ in if prfx_len >= 0 andalso implode (drop (prfx_len, cs)) = sfx then implode (take (prfx_len, cs)) - else raise LIST "unsuffix" + else raise Fail "unsuffix" end; (*remove prefix*) @@ -830,7 +811,7 @@ let val (prfx', sfx) = splitAt (size prfx, explode s) in if implode prfx' = prfx then implode sfx - else raise LIST "unprefix" + else raise Fail "unprefix" end; fun replicate_string 0 _ = "" @@ -1242,12 +1223,7 @@ (*Partition list into elements that satisfy predicate and those that don't. Preserves order of elements in both lists.*) -fun partition (pred: 'a->bool) (ys: 'a list) : ('a list * 'a list) = - let fun part ([], answer) = answer - | part (x::xs, (ys, ns)) = if pred(x) - then part (xs, (x::ys, ns)) - else part (xs, (ys, x::ns)) - in part (rev ys, ([], [])) end; +val partition = List.partition; fun partition_eq (eq:'a * 'a -> bool) = @@ -1263,7 +1239,7 @@ let fun part k xs = if k>j then (case xs of [] => [] - | _ => raise LIST "partition_list") + | _ => raise Fail "partition_list") else let val (ns, rest) = partition (p k) xs; in ns :: part(k+1)rest end @@ -1310,4 +1286,5 @@ end; -open Library; +structure OpenLibrary: LIBRARY = Library; +open OpenLibrary; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/meta_simplifier.ML --- a/src/Pure/meta_simplifier.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/meta_simplifier.ML Thu Mar 03 12:43:01 2005 +0100 @@ -402,7 +402,7 @@ all its Vars? Better: a dynamic check each time a rule is applied. *) fun rewrite_rule_extra_vars prems elhs erhs = - not (term_varnames erhs subset foldl add_term_varnames (term_varnames elhs, prems)) + not (term_varnames erhs subset Library.foldl add_term_varnames (term_varnames elhs, prems)) orelse not (term_tvars erhs subset (term_tvars elhs union List.concat (map term_tvars prems))); @@ -487,16 +487,16 @@ end; fun extract_rews (Simpset (_, {mk_rews = {mk, ...}, ...}), thms) = - flat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk thm)) thms); + List.concat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk thm)) thms); fun orient_comb_simps comb mk_rrule (ss, thms) = let val rews = extract_rews (ss, thms); - val rrules = flat (map mk_rrule rews); - in foldl comb (ss, rrules) end; + val rrules = List.concat (map mk_rrule rews); + in Library.foldl comb (ss, rrules) end; fun extract_safe_rrules (ss, thm) = - flat (map (orient_rrule ss) (extract_rews (ss, [thm]))); + List.concat (map (orient_rrule ss) (extract_rews (ss, [thm]))); (*add rewrite rules explicitly; do not reorient!*) fun ss addsimps thms = @@ -551,7 +551,7 @@ val (lhs, _) = Drule.dest_equals (Drule.strip_imp_concl (Thm.cprop_of thm)) handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm); (*val lhs = Pattern.eta_contract lhs;*) - val a = the (cong_name (head_of (term_of lhs))) handle Option => + val a = valOf (cong_name (head_of (term_of lhs))) handle Option => raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm); val (alist, weak) = congs; val alist2 = overwrite_warn (alist, (a, {lhs = lhs, thm = thm})) @@ -565,11 +565,11 @@ val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm); (*val lhs = Pattern.eta_contract lhs;*) - val a = the (cong_name (head_of lhs)) handle Option => + val a = valOf (cong_name (head_of lhs)) handle Option => raise SIMPLIFIER ("Congruence must start with a constant", thm); val (alist, _) = congs; - val alist2 = filter (fn (x, _) => x <> a) alist; - val weak2 = alist2 |> mapfilter (fn (a, {thm, ...}) => + val alist2 = List.filter (fn (x, _) => x <> a) alist; + val weak2 = alist2 |> List.mapPartial (fn (a, {thm, ...}) => if is_full_cong thm then NONE else SOME a); in ((alist2, weak2), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); @@ -577,8 +577,8 @@ in -val (op addeqcongs) = foldl add_cong; -val (op deleqcongs) = foldl del_cong; +val (op addeqcongs) = Library.foldl add_cong; +val (op deleqcongs) = Library.foldl del_cong; fun ss addcongs congs = ss addeqcongs map (mk_cong ss) congs; fun ss delcongs congs = ss deleqcongs map (mk_cong ss) congs; @@ -607,8 +607,8 @@ in -val (op addsimprocs) = foldl (fn (ss, Simproc procs) => foldl add_proc (ss, procs)); -val (op delsimprocs) = foldl (fn (ss, Simproc procs) => foldl del_proc (ss, procs)); +val (op addsimprocs) = Library.foldl (fn (ss, Simproc procs) => Library.foldl add_proc (ss, procs)); +val (op delsimprocs) = Library.foldl (fn (ss, Simproc procs) => Library.foldl del_proc (ss, procs)); end; @@ -966,7 +966,7 @@ may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*) (let val thm = congc (prover ss, sign, maxidx) cong t0; - val t = if_none (apsome rhs_of thm) t0; + val t = getOpt (Option.map rhs_of thm, t0); val (cl, cr) = Thm.dest_comb t val dVar = Var(("", 0), dummyT) val skel = @@ -993,7 +993,7 @@ in (extract_safe_rrules (ss, asm), SOME asm) end and add_rrules (rrss, asms) ss = - foldl (insert_rrule true) (ss, flat rrss) |> add_prems (mapfilter I asms) + Library.foldl (insert_rrule true) (ss, List.concat rrss) |> add_prems (List.mapPartial I asms) and disch r (prem, eq) = let @@ -1018,12 +1018,12 @@ let val ss' = add_rrules (rev rrss, rev asms) ss; val concl' = - Drule.mk_implies (prem, if_none (apsome rhs_of eq) concl); - val dprem = apsome (curry (disch false) prem) + Drule.mk_implies (prem, getOpt (Option.map rhs_of eq, concl)); + val dprem = Option.map (curry (disch false) prem) in case rewritec (prover, sign, maxidx) ss' concl' of NONE => rebuild prems concl' rrss asms ss (dprem eq) - | SOME (eq', _) => transitive2 (foldl (disch false o swap) - (the (transitive3 (dprem eq) eq'), prems)) + | SOME (eq', _) => transitive2 (Library.foldl (disch false o swap) + (valOf (transitive3 (dprem eq) eq'), prems)) (mut_impc0 (rev prems) (rhs_of eq') (rev rrss) (rev asms) ss) end @@ -1036,8 +1036,8 @@ end and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k = - transitive1 (foldl (fn (eq2, (eq1, prem)) => transitive1 eq1 - (apsome (curry (disch false) prem) eq2)) (NONE, eqns ~~ prems')) + transitive1 (Library.foldl (fn (eq2, (eq1, prem)) => transitive1 eq1 + (Option.map (curry (disch false) prem) eq2)) (NONE, eqns ~~ prems')) (if changed > 0 then mut_impc (rev prems') concl (rev rrss') (rev asms') [] [] [] [] ss ~1 changed @@ -1055,20 +1055,20 @@ let val prem' = rhs_of eqn; val tprems = map term_of prems; - val i = 1 + foldl Int.max (~1, map (fn p => + val i = 1 + Library.foldl Int.max (~1, map (fn p => find_index_eq p tprems) (#hyps (rep_thm eqn))); val (rrs', asm') = rules_of_prem ss prem' in mut_impc prems concl rrss asms (prem' :: prems') - (rrs' :: rrss') (asm' :: asms') (SOME (foldr (disch true) - (take (i, prems), Drule.imp_cong' eqn (reflexive (Drule.list_implies - (drop (i, prems), concl))))) :: eqns) ss (length prems') ~1 + (rrs' :: rrss') (asm' :: asms') (SOME (Library.foldr (disch true) + (Library.take (i, prems), Drule.imp_cong' eqn (reflexive (Drule.list_implies + (Library.drop (i, prems), concl))))) :: eqns) ss (length prems') ~1 end (*legacy code - only for backwards compatibility*) and nonmut_impc ct ss = let val (prem, conc) = dest_implies ct; val thm1 = if simprem then botc skel0 ss prem else NONE; - val prem1 = if_none (apsome rhs_of thm1) prem; + val prem1 = getOpt (Option.map rhs_of thm1, prem); val ss1 = if not useprem then ss else add_rrules (apsnd single (apfst single (rules_of_prem ss prem1))) ss in (case botc skel0 ss1 conc of @@ -1169,7 +1169,7 @@ let val Simpset (_, {solvers = (unsafe_solvers, _), ...}) = ss; val tacf = solve_all_tac unsafe_solvers; - fun prover s th = apsome #1 (Seq.pull (tacf s th)); + fun prover s th = Option.map #1 (Seq.pull (tacf s th)); in rew mode prover ss thm end; val simp_thm = simp rewrite_thm; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/net.ML --- a/src/Pure/net.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/net.ML Thu Mar 03 12:43:01 2005 +0100 @@ -165,9 +165,9 @@ (*Skipping a term in a net. Recursively skip 2 levels if a combination*) fun net_skip (Leaf _, nets) = nets | net_skip (Net{comb,var,alist}, nets) = - foldr net_skip + Library.foldr net_skip (net_skip (comb,[]), - foldr (fn ((_,net), nets) => net::nets) (alist, var::nets)); + Library.foldr (fn ((_,net), nets) => net::nets) (alist, var::nets)); (** Matching and Unification**) @@ -185,7 +185,7 @@ let fun rands _ (Leaf _, nets) = nets | rands t (Net{comb,alist,...}, nets) = case t of - f$t => foldr (matching unif t) (rands f (comb,[]), nets) + f$t => Library.foldr (matching unif t) (rands f (comb,[]), nets) | Const(c,_) => look1 (alist, c) nets | Free(c,_) => look1 (alist, c) nets | Bound i => look1 (alist, string_of_bound i) nets @@ -223,7 +223,7 @@ | dest (Net {comb, var, alist}) = map (cons_fst CombK) (dest comb) @ map (cons_fst VarK) (dest var) @ - flat (map (fn (a, net) => map (cons_fst (AtomK a)) (dest net)) alist); + List.concat (map (fn (a, net) => map (cons_fst (AtomK a)) (dest net)) alist); (** merge **) @@ -232,7 +232,7 @@ insert (keys_x, net, eq) handle INSERT => net; fun merge (net1, net2, eq) = - foldl (add eq) (net1, dest net2); + Library.foldl (add eq) (net1, dest net2); end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/pattern.ML --- a/src/Pure/pattern.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/pattern.ML Thu Mar 03 12:43:01 2005 +0100 @@ -48,7 +48,7 @@ fun string_of_term sg env binders t = Sign.string_of_term sg (Envir.norm_term env (subst_bounds(map Free binders,t))); -fun bname binders i = fst(nth_elem(i,binders)); +fun bname binders i = fst(List.nth(binders,i)); fun bnames binders is = space_implode " " (map (bname binders) is); fun typ_clash sg (tye,T,U) = @@ -113,10 +113,10 @@ fun idx [] j = raise Unif | idx(i::is) j = if i=j then length is else idx is j; -fun at xs i = nth_elem (i,xs); +fun at xs i = List.nth (xs,i); fun mkabs (binders,is,t) = - let fun mk(i::is) = let val (x,T) = nth_elem(i,binders) + let fun mk(i::is) = let val (x,T) = List.nth(binders,i) in Abs(x,T,mk is) end | mk [] = t in mk is end; @@ -158,7 +158,7 @@ (* mk_proj_list(is) = [ |is| - k | 1 <= k <= |is| and is[k] >= 0 ] *) fun mk_proj_list is = - let fun mk(i::is,j) = if is_some i then j :: mk(is,j-1) else mk(is,j-1) + let fun mk(i::is,j) = if isSome i then j :: mk(is,j-1) else mk(is,j-1) | mk([],_) = [] in mk(is,length is - 1) end; @@ -182,7 +182,7 @@ let val js = ints_of' env ts; val js' = map (try (trans d)) js; val ks = mk_proj_list js'; - val ls = mapfilter I js' + val ls = List.mapPartial I js' val Hty = type_of_G env (Fty,length js,ks) val (env',H) = Envir.genvar a (env,Hty) val env'' = @@ -260,11 +260,11 @@ and rigidrigid sg (env,binders,(a,Ta),(b,Tb),ss,ts) = if a<>b then (clash a b; raise Unif) - else foldl (unif sg binders) (unify_types sg (Ta,Tb,env), ss~~ts) + else Library.foldl (unif sg binders) (unify_types sg (Ta,Tb,env), ss~~ts) and rigidrigidB sg (env,binders,i,j,ss,ts) = if i <> j then (clashBB binders i j; raise Unif) - else foldl (unif sg binders) (env ,ss~~ts) + else Library.foldl (unif sg binders) (env ,ss~~ts) and flexrigid sg (params as (env,binders,F,is,t)) = if occurs(F,t,env) then (ocheck_fail sg (F,t,binders,env); raise Unif) @@ -272,7 +272,7 @@ in Envir.update((F,mkabs(binders,is,u)),env') end handle Unif => (proj_fail sg params; raise Unif)); -fun unify(sg,env,tus) = foldl (unif sg []) (env,tus); +fun unify(sg,env,tus) = Library.foldl (unif sg []) (env,tus); (*Eta-contract a term (fully)*) @@ -407,7 +407,7 @@ and cases(binders,env as (iTs,itms),pat,obj) = let val (ph,pargs) = strip_comb pat fun rigrig1(iTs,oargs) = - foldl (mtch binders) ((iTs,itms), pargs~~oargs) + Library.foldl (mtch binders) ((iTs,itms), pargs~~oargs) fun rigrig2((a,Ta),(b,Tb),oargs) = if a<> b then raise MATCH else rigrig1(typ_match tsg (iTs,(Ta,Tb)), oargs) @@ -476,7 +476,7 @@ val skel0 = Bound 0; val rhs_names = - foldr (fn ((_, rhs), names) => add_term_free_names (rhs, names)) (rules, []); + Library.foldr (fn ((_, rhs), names) => add_term_free_names (rhs, names)) (rules, []); fun variant_absfree (x, T, t) = let @@ -485,29 +485,29 @@ in (fn u => Abs (x, T, abstract_over (Free (x', T), u)), t') end; fun match_rew tm (tm1, tm2) = - let val rtm = if_none (Term.rename_abs tm1 tm tm2) tm2 + let val rtm = getOpt (Term.rename_abs tm1 tm tm2, tm2) in SOME (subst_vars (match tsig (tm1, tm)) rtm, rtm) handle MATCH => NONE end; fun rew (Abs (_, _, body) $ t) = SOME (subst_bound (t, body), skel0) | rew tm = (case get_first (match_rew tm) rules of - NONE => apsome (rpair skel0) (get_first (fn p => p tm) procs) + NONE => Option.map (rpair skel0) (get_first (fn p => p tm) procs) | x => x); fun rew1 (Var _) _ = NONE | rew1 skel tm = (case rew2 skel tm of SOME tm1 => (case rew tm1 of - SOME (tm2, skel') => SOME (if_none (rew1 skel' tm2) tm2) + SOME (tm2, skel') => SOME (getOpt (rew1 skel' tm2, tm2)) | NONE => SOME tm1) | NONE => (case rew tm of - SOME (tm1, skel') => SOME (if_none (rew1 skel' tm1) tm1) + SOME (tm1, skel') => SOME (getOpt (rew1 skel' tm1, tm1)) | NONE => NONE)) and rew2 skel (tm1 $ tm2) = (case tm1 of Abs (_, _, body) => let val tm' = subst_bound (tm2, body) - in SOME (if_none (rew2 skel0 tm') tm') end + in SOME (getOpt (rew2 skel0 tm', tm')) end | _ => let val (skel1, skel2) = (case skel of skel1 $ skel2 => (skel1, skel2) @@ -530,7 +530,7 @@ end | rew2 _ _ = NONE - in if_none (rew1 skel0 tm) tm end; + in getOpt (rew1 skel0 tm, tm) end; end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/proof_general.ML --- a/src/Pure/proof_general.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/proof_general.ML Thu Mar 03 12:43:01 2005 +0100 @@ -179,8 +179,8 @@ ([("class", pgip_class), ("origin", pgip_origin), ("id", pgip_id)] @ - (if_none (apsome (single o (pair "refseq")) (!pgip_refseq)) []) @ - (if_none (apsome (single o (pair "refid")) (!pgip_refid)) []) @ + getOpt (Option.map (single o (pair "refseq")) (!pgip_refseq), []) @ + getOpt (Option.map (single o (pair "refid")) (!pgip_refid), []) @ [("seq", string_of_int (Library.inc myseq_no))]) pgips in @@ -395,14 +395,14 @@ fun trace_action action name = if action = ThyInfo.Update then - seq tell_file_loaded (ThyInfo.loaded_files name) + List.app tell_file_loaded (ThyInfo.loaded_files name) else if action = ThyInfo.Outdate orelse action = ThyInfo.Remove then - seq tell_file_retracted (add_master_files name (ThyInfo.loaded_files name)) + List.app tell_file_retracted (add_master_files name (ThyInfo.loaded_files name)) else (); in fun setup_thy_loader () = ThyInfo.add_hook trace_action; - fun sync_thy_loader () = seq (trace_action ThyInfo.Update) (ThyInfo.names ()); + fun sync_thy_loader () = List.app (trace_action ThyInfo.Update) (ThyInfo.names ()); end; @@ -421,7 +421,7 @@ fun try_update_thy_only file = ThyLoad.cond_add_path (Path.dir (Path.unpack file)) (fn () => let val name = thy_name file in - if is_some (ThyLoad.check_file NONE (ThyLoad.thy_path name)) + if isSome (ThyLoad.check_file NONE (ThyLoad.thy_path name)) then update_thy_only name else warning ("Unkown theory context of ML file." ^ which_context ()) end) (); @@ -492,7 +492,7 @@ conditional (thm_depsN mem_string ! print_mode) (fn () => let val names = map Thm.name_of_thm ths \ ""; - val deps = Symtab.keys (foldl (uncurry Proofterm.thms_of_proof) + val deps = Symtab.keys (Library.foldl (uncurry Proofterm.thms_of_proof) (Symtab.empty, map Thm.proof_of ths)) \ ""; in if null names orelse null deps then () @@ -502,7 +502,7 @@ in fun setup_present_hook () = - Present.add_hook (fn _ => fn res => tell_thm_deps (flat (map #2 res))); + Present.add_hook (fn _ => fn res => tell_thm_deps (List.concat (map #2 res))); end; @@ -785,7 +785,7 @@ case (!keywords_classification_table) of SOME t => t | NONE => (let - val tab = foldl (fn (tab,(c,_,k,_))=>Symtab.update((c,k),tab)) + val tab = Library.foldl (fn (tab,(c,_,k,_))=>Symtab.update((c,k),tab)) (Symtab.empty,OuterSyntax.dest_parsers()) in (keywords_classification_table := SOME tab; tab) end) @@ -969,10 +969,10 @@ ((markup_comment_whs prewhs) @ (if (length rest2 = length rest1) then [] else markup_text (implode - (map OuterLex.unparse (take (length rest1 - length rest2, rest1)))) + (map OuterLex.unparse (Library.take (length rest1 - length rest2, rest1)))) "litcomment") @ (markup_comment_whs postwhs), - take (length rest3 - 1,rest3)) + Library.take (length rest3 - 1,rest3)) end fun xmls_of_impropers toks = @@ -1051,7 +1051,7 @@ issue_pgipe "usespgml" [("version", isabelle_pgml_version_supported)]; fun hasprefs () = - seq (fn (prefcat, prefs) => + List.app (fn (prefcat, prefs) => issue_pgips [XML.element "hasprefs" [("prefcategory",prefcat)] (map (fn (name, (descr, (ty, (default,_),_))) => XML.element "haspref" [("name", name), @@ -1059,7 +1059,7 @@ ("default", default)] [ty]) prefs)]) (!preferences); -fun allprefs () = foldl (op @) ([], map snd (!preferences)) +fun allprefs () = Library.foldl (op @) ([], map snd (!preferences)) fun setpref name value = (case assoc (allprefs(), name) of @@ -1225,7 +1225,7 @@ pgip_refid := xmlattro "id" attrs) in if (class = "pa") then - seq process_pgip_element pgips + List.app process_pgip_element pgips else raise PGIP "PGIP packet for me should have class=\"pa\"" end) @@ -1387,7 +1387,7 @@ "\n '(" ^ space_implode "\n " (map (quote o regexp_quote) strs) ^ "))\n"; fun make_elisp_commands commands kind = - defconst kind (mapfilter + defconst kind (List.mapPartial (fn (c, _, k, _) => if k = kind then SOME c else NONE) commands); @@ -1399,7 +1399,7 @@ \;; $" ^ "Id$\n\ \;;\n" ^ defconst "major" (map #1 commands) ^ - defconst "minor" (filter Syntax.is_identifier keywords) ^ + defconst "minor" (List.filter Syntax.is_identifier keywords) ^ implode (map (make_elisp_commands commands) OuterSyntax.Keyword.kinds) ^ "\n(provide 'isar-keywords)\n"; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/proofterm.ML Thu Mar 03 12:43:01 2005 +0100 @@ -140,7 +140,7 @@ | SOME ps => if prop mem ps then tabs else oras_of ((Symtab.update ((name, prop::ps), thms), oras), prf)) | oras_of ((thms, oras), prf as Oracle _) = (thms, prf ins oras) - | oras_of (tabs, MinProof prfs) = foldl oras_of (tabs, prfs) + | oras_of (tabs, MinProof prfs) = Library.foldl oras_of (tabs, prfs) | oras_of (tabs, _) = tabs in snd (oras_of ((Symtab.empty, prfs), prf)) @@ -155,7 +155,7 @@ NONE => thms_of_proof (Symtab.update ((s, [(prop, prf')]), tab)) prf | SOME ps => if exists (equal prop o fst) ps then tab else thms_of_proof (Symtab.update ((s, (prop, prf')::ps), tab)) prf) - | thms_of_proof tab (MinProof prfs) = foldl (uncurry thms_of_proof) (tab, prfs) + | thms_of_proof tab (MinProof prfs) = Library.foldl (uncurry thms_of_proof) (tab, prfs) | thms_of_proof tab _ = tab; fun axms_of_proof tab (Abst (_, _, prf)) = axms_of_proof tab prf @@ -163,7 +163,7 @@ | axms_of_proof tab (prf1 %% prf2) = axms_of_proof (axms_of_proof tab prf1) prf2 | axms_of_proof tab (prf % _) = axms_of_proof tab prf | axms_of_proof tab (prf as PAxm (s, _, _)) = Symtab.update ((s, prf), tab) - | axms_of_proof tab (MinProof prfs) = foldl (uncurry axms_of_proof) (tab, prfs) + | axms_of_proof tab (MinProof prfs) = Library.foldl (uncurry axms_of_proof) (tab, prfs) | axms_of_proof tab _ = tab; (** collect all theorems, axioms and oracles **) @@ -199,9 +199,9 @@ fun (prf %> t) = prf % SOME t; -val proof_combt = foldl (op %>); -val proof_combt' = foldl (op %); -val proof_combP = foldl (op %%); +val proof_combt = Library.foldl (op %>); +val proof_combt' = Library.foldl (op %); +val proof_combP = Library.foldl (op %%); fun strip_combt prf = let fun stripc (prf % t, ts) = stripc (prf, t::ts) @@ -217,7 +217,7 @@ (PThm (_, prf', _, _), _) => prf' | _ => prf); -val mk_Abst = foldr (fn ((s, T:typ), prf) => Abst (s, NONE, prf)); +val mk_Abst = Library.foldr (fn ((s, T:typ), prf) => Abst (s, NONE, prf)); fun mk_AbsP (i, prf) = funpow i (fn prf => AbsP ("H", NONE, prf)) prf; fun apsome' f NONE = raise SAME @@ -233,7 +233,7 @@ handle SAME => Abst (s, T, mapp prf)) | mapp (AbsP (s, t, prf)) = (AbsP (s, apsome' (same f) t, mapph prf) handle SAME => AbsP (s, t, mapp prf)) - | mapp (prf % t) = (mapp prf % apsome f t + | mapp (prf % t) = (mapp prf % Option.map f t handle SAME => prf % apsome' (same f) t) | mapp (prf1 %% prf2) = (mapp prf1 %% mapph prf2 handle SAME => prf1 %% mapp prf2) @@ -254,8 +254,8 @@ | fold_proof_terms f g (a, prf % NONE) = fold_proof_terms f g (a, prf) | fold_proof_terms f g (a, prf1 %% prf2) = fold_proof_terms f g (fold_proof_terms f g (a, prf1), prf2) - | fold_proof_terms _ g (a, PThm (_, _, _, SOME Ts)) = foldr g (Ts, a) - | fold_proof_terms _ g (a, PAxm (_, prop, SOME Ts)) = foldr g (Ts, a) + | fold_proof_terms _ g (a, PThm (_, _, _, SOME Ts)) = Library.foldr g (Ts, a) + | fold_proof_terms _ g (a, PAxm (_, prop, SOME Ts)) = Library.foldr g (Ts, a) | fold_proof_terms _ _ (a, _) = a; val add_prf_names = fold_proof_terms add_term_names ((uncurry K) o swap); @@ -282,7 +282,7 @@ fun strip_abs (_::Ts) (Abs (_, _, t)) = strip_abs Ts t | strip_abs _ t = t; -fun mk_abs Ts t = foldl (fn (t', T) => Abs ("", T, t')) (t, Ts); +fun mk_abs Ts t = Library.foldl (fn (t', T) => Abs ("", T, t')) (t, Ts); (*Abstraction of a proof term over its occurrences of v, @@ -304,7 +304,7 @@ | abst lev (Abst (a, T, prf)) = Abst (a, T, abst (lev + 1) prf) | abst lev (prf1 %% prf2) = (abst lev prf1 %% absth lev prf2 handle SAME => prf1 %% abst lev prf2) - | abst lev (prf % t) = (abst lev prf % apsome (absth' lev) t + | abst lev (prf % t) = (abst lev prf % Option.map (absth' lev) t handle SAME => prf % apsome' (abst' lev) t) | abst _ _ = raise SAME and absth lev prf = (abst lev prf handle SAME => prf) @@ -331,7 +331,7 @@ (prf_incr_bv' incP inct Plev tlev prf %% prf_incr_bv incP inct Plev tlev prf' handle SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') | prf_incr_bv' incP inct Plev tlev (prf % t) = - (prf_incr_bv' incP inct Plev tlev prf % apsome (incr_bv' inct tlev) t + (prf_incr_bv' incP inct Plev tlev prf % Option.map (incr_bv' inct tlev) t handle SAME => prf % apsome' (same (incr_bv' inct tlev)) t) | prf_incr_bv' _ _ _ _ _ = raise SAME and prf_incr_bv incP inct Plev tlev prf = @@ -383,7 +383,7 @@ handle SAME => Abst (s, T, norm prf)) | norm (AbsP (s, t, prf)) = (AbsP (s, apsome' (norm_term_same env) t, normh prf) handle SAME => AbsP (s, t, norm prf)) - | norm (prf % t) = (norm prf % apsome (norm_term env) t + | norm (prf % t) = (norm prf % Option.map (norm_term env) t handle SAME => prf % apsome' (norm_term_same env) t) | norm (prf1 %% prf2) = (norm prf1 %% normh prf2 handle SAME => prf1 %% norm prf2) @@ -425,7 +425,7 @@ | subst lev (Abst (a, T, body)) = Abst (a, T, subst (lev+1) body) | subst lev (prf %% prf') = (subst lev prf %% substh lev prf' handle SAME => prf %% subst lev prf') - | subst lev (prf % t) = (subst lev prf % apsome (substh' lev) t + | subst lev (prf % t) = (subst lev prf % Option.map (substh' lev) t handle SAME => prf % apsome' (subst' lev) t) | subst _ _ = raise SAME and substh lev prf = (subst lev prf handle SAME => prf) @@ -451,7 +451,7 @@ (**** Freezing and thawing of variables in proof terms ****) fun frzT names = - map_type_tvar (fn (ixn, xs) => TFree (the (assoc (names, ixn)), xs)); + map_type_tvar (fn (ixn, xs) => TFree (valOf (assoc (names, ixn)), xs)); fun thawT names = map_type_tfree (fn (s, xs) => case assoc (names, s) of @@ -465,7 +465,7 @@ | freeze names names' (Const (s, T)) = Const (s, frzT names' T) | freeze names names' (Free (s, T)) = Free (s, frzT names' T) | freeze names names' (Var (ixn, T)) = - Free (the (assoc (names, ixn)), frzT names' T) + Free (valOf (assoc (names, ixn)), frzT names' T) | freeze names names' t = t; fun thaw names names' (t $ u) = @@ -557,7 +557,7 @@ let val used = it_term_types add_typ_tfree_names (t, []) and tvars = map #1 (it_term_types add_typ_tvars (t, [])); - val (alist, _) = foldr new_name (tvars, ([], used)); + val (alist, _) = Library.foldr new_name (tvars, ([], used)); in (case alist of [] => prf (*nothing to do!*) @@ -615,7 +615,7 @@ | lift' Us Ts (AbsP (s, t, prf)) = (AbsP (s, apsome' (same (lift'' Us Ts)) t, lifth' Us Ts prf) handle SAME => AbsP (s, t, lift' Us Ts prf)) - | lift' Us Ts (prf % t) = (lift' Us Ts prf % apsome (lift'' Us Ts) t + | lift' Us Ts (prf % t) = (lift' Us Ts prf % Option.map (lift'' Us Ts) t handle SAME => prf % apsome' (same (lift'' Us Ts)) t) | lift' Us Ts (prf1 %% prf2) = (lift' Us Ts prf1 %% lifth' Us Ts prf2 handle SAME => prf1 %% lift' Us Ts prf2) @@ -637,7 +637,7 @@ | lift Us bs i j (Const ("all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), lift (T::Us) (false::bs) i (j+1) t) | lift Us bs i j _ = proof_combP (lifth' (rev Us) [] prf, - map (fn k => (#3 (foldr mk_app (bs, (i-1, j-1, PBound k))))) + map (fn k => (#3 (Library.foldr mk_app (bs, (i-1, j-1, PBound k))))) (i + k - 1 downto i)); in mk_AbsP (k, lift [] [] 0 0 Bi) @@ -671,7 +671,7 @@ in mk_AbsP (lb+la, proof_combP (sprf, map PBound (lb + la - 1 downto la)) %% - proof_combP (rprf, (if n>0 then [mk_asm_prf (the A) (~n)] else []) @ + proof_combP (rprf, (if n>0 then [mk_asm_prf (valOf A) (~n)] else []) @ map (flatten_params_proof 0 0 n) (oldAs ~~ (la - 1 downto 0)))) end; @@ -726,7 +726,7 @@ abstract_rule_axm % NONE % NONE %% forall_intr_proof x a prf; fun check_comb (PAxm ("ProtoPure.combination", _, _) % f % g % _ % _ %% prf %% _) = - is_some f orelse check_comb prf + isSome f orelse check_comb prf | check_comb (PAxm ("ProtoPure.transitive", _, _) % _ % _ % _ %% prf1 %% prf2) = check_comb prf1 andalso check_comb prf2 | check_comb (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf) = check_comb prf @@ -779,7 +779,7 @@ fun add_funvars Ts (vs, t) = if is_fun (fastype_of1 (Ts, t)) then - vs union mapfilter (fn Var (ixn, T) => + vs union List.mapPartial (fn Var (ixn, T) => if is_fun T then SOME ixn else NONE | _ => NONE) (vars_of t) else vs; @@ -791,10 +791,10 @@ | add_npvars _ _ Ts (vs, t) = add_npvars' Ts (vs, t) and add_npvars' Ts (vs, t) = (case strip_comb t of (Var (ixn, _), ts) => if test_args [] ts then vs - else foldl (add_npvars' Ts) (overwrite (vs, - (ixn, foldl (add_funvars Ts) (if_none (assoc (vs, ixn)) [], ts))), ts) - | (Abs (_, T, u), ts) => foldl (add_npvars' (T::Ts)) (vs, u :: ts) - | (_, ts) => foldl (add_npvars' Ts) (vs, ts)); + else Library.foldl (add_npvars' Ts) (overwrite (vs, + (ixn, Library.foldl (add_funvars Ts) (getOpt (assoc (vs, ixn), []), ts))), ts) + | (Abs (_, T, u), ts) => Library.foldl (add_npvars' (T::Ts)) (vs, u :: ts) + | (_, ts) => Library.foldl (add_npvars' Ts) (vs, ts)); fun prop_vars (Const ("==>", _) $ P $ Q) = prop_vars P union prop_vars Q | prop_vars (Const ("all", _) $ Abs (_, _, t)) = prop_vars t @@ -814,7 +814,7 @@ end; fun needed_vars prop = - foldl op union ([], map op ins (add_npvars true true [] ([], prop))) union + Library.foldl op union ([], map op ins (add_npvars true true [] ([], prop))) union prop_vars prop; fun gen_axm_proof c name prop = @@ -835,7 +835,7 @@ in (b, is, ch, if ch then Abst (a, T, body') else prf) end | shrink ls lev (prf as AbsP (a, t, body)) = let val (b, is, ch, body') = shrink (lev::ls) lev body - in (b orelse 0 mem is, mapfilter (fn 0 => NONE | i => SOME (i-1)) is, + in (b orelse 0 mem is, List.mapPartial (fn 0 => NONE | i => SOME (i-1)) is, ch, if ch then AbsP (a, t, body') else prf) end | shrink ls lev prf = @@ -852,9 +852,9 @@ let val (is, ch, (ch', t')::ts', prf') = shrink' ls lev (t::ts) prfs prf1 in (is, ch orelse ch', ts', if ch orelse ch' then prf' % t' else prf) end | shrink' ls lev ts prfs (prf as PBound i) = - (if exists (fn SOME (Bound j) => lev-j <= nth_elem (i, ls) | _ => true) ts + (if exists (fn SOME (Bound j) => lev-j <= List.nth (ls, i) | _ => true) ts orelse not (null (duplicates - (foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)))) + (Library.foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)))) orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf) | shrink' ls lev ts prfs (prf as Hyp _) = ([], false, map (pair false) ts, prf) | shrink' ls lev ts prfs (prf as MinProof _) = @@ -865,8 +865,8 @@ | Oracle (_, prop, _) => prop | _ => error "shrink: proof not in normal form"); val vs = vars_of prop; val (ts', ts'') = splitAt (length vs, ts) - val insts = take (length ts', map (fst o dest_Var) vs) ~~ ts'; - val nvs = foldl (fn (ixns', (ixn, ixns)) => + val insts = Library.take (length ts', map (fst o dest_Var) vs) ~~ ts'; + val nvs = Library.foldl (fn (ixns', (ixn, ixns)) => ixn ins (case assoc (insts, ixn) of SOME (SOME t) => if is_proj t then ixns union ixns' else ixns' | _ => ixns union ixns')) @@ -887,7 +887,7 @@ (** see pattern.ML **) -fun flt (i: int) = filter (fn n => n < i); +fun flt (i: int) = List.filter (fn n => n < i); fun fomatch Ts tymatch j = let @@ -916,7 +916,7 @@ (pinst, fomatch Ts tymatch j tinst (t, Envir.beta_norm u)); fun matchT (pinst, (tyinsts, insts)) p = (pinst, (tymatch (tyinsts, K p), insts)); - fun matchTs inst (Ts, Us) = foldl (uncurry matchT) (inst, Ts ~~ Us); + fun matchTs inst (Ts, Us) = Library.foldl (uncurry matchT) (inst, Ts ~~ Us); fun mtch Ts i j (pinst, tinst) (Hyp (Var (ixn, _)), prf) = if i = 0 andalso j = 0 then ((ixn, prf) :: pinst, tinst) @@ -932,10 +932,10 @@ | mtch Ts i j inst (prf1 %% prf2, prf1' %% prf2') = mtch Ts i j (mtch Ts i j inst (prf1, prf1')) (prf2, prf2') | mtch Ts i j inst (Abst (_, opT, prf1), Abst (_, opU, prf2)) = - mtch (if_none opU dummyT :: Ts) i (j+1) + mtch (getOpt (opU,dummyT) :: Ts) i (j+1) (optmatch matchT inst (opT, opU)) (prf1, prf2) | mtch Ts i j inst (prf1, Abst (_, opU, prf2)) = - mtch (if_none opU dummyT :: Ts) i (j+1) inst + mtch (getOpt (opU,dummyT) :: Ts) i (j+1) inst (incr_pboundvars 0 1 prf1 %> Bound 0, prf2) | mtch Ts i j inst (AbsP (_, opt, prf1), AbsP (_, opu, prf2)) = mtch Ts (i+1) j (optmatch (matcht Ts j) inst (opt, opu)) (prf1, prf2) @@ -967,18 +967,18 @@ | subst' _ t = t; fun subst plev tlev (AbsP (a, t, body)) = - AbsP (a, apsome (subst' tlev) t, subst (plev+1) tlev body) + AbsP (a, Option.map (subst' tlev) t, subst (plev+1) tlev body) | subst plev tlev (Abst (a, T, body)) = - Abst (a, apsome substT T, subst plev (tlev+1) body) + Abst (a, Option.map substT T, subst plev (tlev+1) body) | subst plev tlev (prf %% prf') = subst plev tlev prf %% subst plev tlev prf' - | subst plev tlev (prf % t) = subst plev tlev prf % apsome (subst' tlev) t + | subst plev tlev (prf % t) = subst plev tlev prf % Option.map (subst' tlev) t | subst plev tlev (prf as Hyp (Var (ixn, _))) = (case assoc (pinst, ixn) of NONE => prf | SOME prf' => incr_pboundvars plev tlev prf') | subst _ _ (PThm (id, prf, prop, Ts)) = - PThm (id, prf, prop, apsome (map substT) Ts) + PThm (id, prf, prop, Option.map (map substT) Ts) | subst _ _ (PAxm (id, prop, Ts)) = - PAxm (id, prop, apsome (map substT) Ts) + PAxm (id, prop, Option.map (map substT) Ts) | subst _ _ t = t in subst 0 0 end; @@ -1023,42 +1023,42 @@ SOME prf' => SOME (prf', skel0) | NONE => get_first (fn (prf1, prf2) => SOME (prf_subst (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2) - handle PMatch => NONE) (filter (could_unify prf o fst) rules)); + handle PMatch => NONE) (List.filter (could_unify prf o fst) rules)); fun rew0 Ts (prf as AbsP (_, _, prf' %% PBound 0)) = if prf_loose_Pbvar1 prf' 0 then rew Ts prf else let val prf'' = incr_pboundvars (~1) 0 prf' - in SOME (if_none (rew Ts prf'') (prf'', skel0)) end + in SOME (getOpt (rew Ts prf'', (prf'', skel0))) end | rew0 Ts (prf as Abst (_, _, prf' % SOME (Bound 0))) = if prf_loose_bvar1 prf' 0 then rew Ts prf else let val prf'' = incr_pboundvars 0 (~1) prf' - in SOME (if_none (rew Ts prf'') (prf'', skel0)) end + in SOME (getOpt (rew Ts prf'', (prf'', skel0))) end | rew0 Ts prf = rew Ts prf; fun rew1 _ (Hyp (Var _)) _ = NONE | rew1 Ts skel prf = (case rew2 Ts skel prf of SOME prf1 => (case rew0 Ts prf1 of - SOME (prf2, skel') => SOME (if_none (rew1 Ts skel' prf2) prf2) + SOME (prf2, skel') => SOME (getOpt (rew1 Ts skel' prf2, prf2)) | NONE => SOME prf1) | NONE => (case rew0 Ts prf of - SOME (prf1, skel') => SOME (if_none (rew1 Ts skel' prf1) prf1) + SOME (prf1, skel') => SOME (getOpt (rew1 Ts skel' prf1, prf1)) | NONE => NONE)) and rew2 Ts skel (prf % SOME t) = (case prf of Abst (_, _, body) => let val prf' = prf_subst_bounds [t] body - in SOME (if_none (rew2 Ts skel0 prf') prf') end + in SOME (getOpt (rew2 Ts skel0 prf', prf')) end | _ => (case rew1 Ts (case skel of skel' % _ => skel' | _ => skel0) prf of SOME prf' => SOME (prf' % SOME t) | NONE => NONE)) - | rew2 Ts skel (prf % NONE) = apsome (fn prf' => prf' % NONE) + | rew2 Ts skel (prf % NONE) = Option.map (fn prf' => prf' % NONE) (rew1 Ts (case skel of skel' % _ => skel' | _ => skel0) prf) | rew2 Ts skel (prf1 %% prf2) = (case prf1 of AbsP (_, _, body) => let val prf' = prf_subst_pbounds [prf2] body - in SOME (if_none (rew2 Ts skel0 prf') prf') end + in SOME (getOpt (rew2 Ts skel0 prf', prf')) end | _ => let val (skel1, skel2) = (case skel of skel1 %% skel2 => (skel1, skel2) @@ -1071,7 +1071,7 @@ SOME prf2' => SOME (prf1 %% prf2') | NONE => NONE) end) - | rew2 Ts skel (Abst (s, T, prf)) = (case rew1 (if_none T dummyT :: Ts) + | rew2 Ts skel (Abst (s, T, prf)) = (case rew1 (getOpt (T,dummyT) :: Ts) (case skel of Abst (_, _, skel') => skel' | _ => skel0) prf of SOME prf' => SOME (Abst (s, T, prf')) | NONE => NONE) @@ -1081,7 +1081,7 @@ | NONE => NONE) | rew2 _ _ _ = NONE - in if_none (rew1 [] skel0 prf) prf end; + in getOpt (rew1 [] skel0 prf, prf) end; fun rewrite_proof tsig = rewrite_prf (fn (tab, f) => Type.typ_match tsig (tab, f ()) handle Type.TYPE_MATCH => raise PMatch); @@ -1127,7 +1127,7 @@ map SOME (sort (make_ord atless) (term_frees prop)); val opt_prf = if ! proofs = 2 then #4 (shrink [] 0 (rewrite_prf fst (ProofData.get_sg sign) - (foldr (uncurry implies_intr_proof) (hyps, prf)))) + (Library.foldr (uncurry implies_intr_proof) (hyps, prf)))) else MinProof (mk_min_proof ([], prf)); val head = (case strip_combt (fst (strip_combP prf)) of (PThm ((old_name, _), prf', prop', NONE), args') => diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/pure_thy.ML --- a/src/Pure/pure_thy.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/pure_thy.ML Thu Mar 03 12:43:01 2005 +0100 @@ -143,7 +143,7 @@ fun select_thm (s, NONE) xs = xs | select_thm (s, SOME is) xs = map - (fn i => (if i < 1 then raise LIST "" else nth_elem (i-1, xs)) handle LIST _ => + (fn i => (if i < 1 then raise Subscript else List.nth (xs, i-1)) handle Subscript => error ("Bad subscript " ^ string_of_int i ^ " for " ^ quote s)) is; @@ -157,7 +157,7 @@ val ref {space, thms_tab, ...} = get_theorems thy; in fn name => - apsome (map (Thm.transfer_sg (Sign.deref sg_ref))) (*semi-dynamic identity*) + Option.map (map (Thm.transfer_sg (Sign.deref sg_ref))) (*semi-dynamic identity*) (Symtab.lookup (thms_tab, NameSpace.intern space name)) (*static content*) end; @@ -178,7 +178,7 @@ get_first (fn thy => lookup_thms thy name) (theory :: Theory.ancestors_of theory) |> the_thms name |> select_thm namei |> map (Thm.transfer theory); -fun get_thmss thy names = flat (map (get_thms thy) names); +fun get_thmss thy names = List.concat (map (get_thms thy) names); fun get_thm thy (namei as (name, _)) = single_thm name (get_thms thy namei); @@ -186,7 +186,7 @@ fun thms_of thy = let val ref {thms_tab, ...} = get_theorems thy in - map (fn th => (Thm.name_of_thm th, th)) (flat (map snd (Symtab.dest thms_tab))) + map (fn th => (Thm.name_of_thm th, th)) (List.concat (map snd (Symtab.dest thms_tab))) end; @@ -200,12 +200,12 @@ | SOME ths' => Library.equal_lists Thm.eq_thm (ths, ths')); in (thy :: Theory.ancestors_of thy) - |> map (gen_distinct eq_fst o filter valid o FactIndex.find idx o #index o ! o get_theorems) - |> flat + |> map (gen_distinct eq_fst o List.filter valid o FactIndex.find idx o #index o ! o get_theorems) + |> List.concat end; fun thms_containing_consts thy consts = - thms_containing thy (consts, []) |> map #2 |> flat + thms_containing thy (consts, []) |> map #2 |> List.concat |> map (fn th => (Thm.name_of_thm th, th)) (* intro/elim theorems *) @@ -235,7 +235,7 @@ fun substsize prop = let val pat = extract_term prop val (_,subst) = Pattern.match tsig (pat,obj) - in foldl op+ (0, map (size_of_term o snd) subst) end + in Library.foldl op+ (0, map (size_of_term o snd) subst) end fun thm_ord ((p0,s0,_),(p1,s1,_)) = prod_ord (int_ord o pairself (fn 0 => 0 | x => 1)) int_ord ((p0,s0),(p1,s1)); @@ -355,7 +355,7 @@ let fun app (x, (ths, atts)) = Thm.applys_attributes ((x, ths), atts); val (thy', thms) = enter_thms (Theory.sign_of thy) - name_thmss (name_thms false) (apsnd flat o foldl_map app) thy + name_thmss (name_thms false) (apsnd List.concat o foldl_map app) thy (bname, map (fn (ths, atts) => (get thy ths, atts @ more_atts @ [kind_att])) ths_atts); in (thy', (bname, thms)) end; @@ -387,7 +387,7 @@ | gen_smart_store_thms name_thm (name, thms) = let val merge_sg = Sign.merge_refs o apsnd (Sign.self_ref o Thm.sign_of_thm); - val sg_ref = foldl merge_sg (Sign.self_ref (Thm.sign_of_thm (hd thms)), tl thms); + val sg_ref = Library.foldl merge_sg (Sign.self_ref (Thm.sign_of_thm (hd thms)), tl thms); in snd (enter_thms (Sign.deref sg_ref) (name_thm true) (name_thm false) I () (name, thms)) end; @@ -404,7 +404,7 @@ in case prop of Const ("all", _) $ Abs (a, T, _) => let val used = map (fst o fst) - (filter (equal i o snd o fst) (Term.add_vars ([], prop))) + (List.filter (equal i o snd o fst) (Term.add_vars ([], prop))) in forall_elim (cterm_of sign (Var ((variant used a, i), T))) th end | _ => raise THM ("forall_elim_var", i, [th]) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/search.ML --- a/src/Pure/search.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/search.ML Thu Mar 03 12:43:01 2005 +0100 @@ -214,8 +214,8 @@ let val tac = tracify trace_BEST_FIRST tac1 fun pairsize th = (sizef th, th); fun bfs (news,nprf_heap) = - (case partition satp news of - ([],nonsats) => next(foldr ThmHeap.insert + (case List.partition satp news of + ([],nonsats) => next(Library.foldr ThmHeap.insert (map pairsize nonsats, nprf_heap)) | (sats,_) => some_of_list sats) and next nprf_heap = @@ -239,7 +239,7 @@ fun gen_BREADTH_FIRST message satpred (tac:tactic) = let val tacf = Seq.list_of o tac; fun bfs prfs = - (case partition satpred prfs of + (case List.partition satpred prfs of ([],[]) => [] | ([],nonsats) => (message("breadth=" ^ string_of_int(length nonsats)); @@ -275,9 +275,9 @@ let val tf = tracify trace_ASTAR tac1; fun bfs (news,nprfs,level) = let fun cost thm = (level, costf level thm, thm) - in (case partition satp news of + in (case List.partition satp news of ([],nonsats) - => next (foldr insert_with_level (map cost nonsats, nprfs)) + => next (Library.foldr insert_with_level (map cost nonsats, nprfs)) | (sats,_) => some_of_list sats) end and next [] = NONE diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/sign.ML --- a/src/Pure/sign.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/sign.ML Thu Mar 03 12:43:01 2005 +0100 @@ -231,7 +231,7 @@ val tsig_of = #tsig o rep_sg; fun is_logtype sg c = c mem_string Type.logical_types (tsig_of sg); -fun const_type (Sg (_, {consts, ...})) c = apsome #1 (Symtab.lookup (consts, c)); +fun const_type (Sg (_, {consts, ...})) c = Option.map #1 (Symtab.lookup (consts, c)); (* id and self *) @@ -478,7 +478,7 @@ (* declare and retrieve names *) fun space_of spaces kind = - if_none (assoc (spaces, kind)) NameSpace.empty; + getOpt (assoc (spaces, kind), NameSpace.empty); (*input and output of qualified names*) fun intrn spaces kind = NameSpace.intern (space_of spaces kind); @@ -514,8 +514,8 @@ fun mapping f add_xs t = let fun f' x = let val y = f x in if x = y then NONE else SOME (x, y) end; - val table = mapfilter f' (add_xs (t, [])); - fun lookup x = if_none (assoc (table, x)) x; + val table = List.mapPartial f' (add_xs (t, [])); + fun lookup x = getOpt (assoc (table, x), x); in lookup end; (*intern / extern typ*) @@ -568,7 +568,7 @@ val full_name = full o #path o rep_sg; fun full_name_path sg elems = - full (SOME (if_none (#path (rep_sg sg)) [] @ NameSpace.unpack elems)); + full (SOME (getOpt (#path (rep_sg sg), []) @ NameSpace.unpack elems)); end; @@ -607,7 +607,7 @@ Syntax.pretty_sort (syn_of sg) (if ! NameSpace.long_names then S else extrn_sort spaces S); -fun pretty_classrel sg cs = Pretty.block (flat +fun pretty_classrel sg cs = Pretty.block (List.concat (separate [Pretty.str " <", Pretty.brk 1] (map (single o pretty_sort sg o single) cs))); fun pretty_arity sg (t, Ss, S) = @@ -690,11 +690,11 @@ val certify_typ_raw = Type.cert_typ_raw o tsig_of; fun certify_tyname sg c = - if is_some (Symtab.lookup (#types (Type.rep_tsig (tsig_of sg)), c)) then c + if isSome (Symtab.lookup (#types (Type.rep_tsig (tsig_of sg)), c)) then c else raise TYPE ("Undeclared type constructor: " ^ quote c, [], []); fun certify_const sg c = - if is_some (const_type sg c) then c + if isSome (const_type sg c) then c else raise TYPE ("Undeclared constant: " ^ quote c, [], []); @@ -717,7 +717,7 @@ fun typ_of (_, Const (_, T)) = T | typ_of (_, Free (_, T)) = T | typ_of (_, Var (_, T)) = T - | typ_of (bs, Bound i) = snd (nth_elem (i, bs) handle LIST _ => + | typ_of (bs, Bound i) = snd (List.nth (bs, i) handle Subscript => raise TYPE ("Loose bound variable: B." ^ string_of_int i, [], [Bound i])) | typ_of (bs, Abs (x, T, body)) = T --> typ_of ((x, T) :: bs, body) | typ_of (bs, t $ u) = @@ -832,7 +832,7 @@ let val tsig = tsig_of sg; - val termss = foldr multiply (map fst args, [[]]); + val termss = Library.foldr multiply (map fst args, [[]]); val typs = map (fn (_, T) => certify_typ sg T handle TYPE (msg, _, _) => error msg) args; @@ -842,8 +842,8 @@ handle TYPE (msg, _, _) => Error msg; val err_results = map infer termss; - val errs = mapfilter get_error err_results; - val results = mapfilter get_ok err_results; + val errs = List.mapPartial get_error err_results; + val results = List.mapPartial get_ok err_results; val ambiguity = length termss; @@ -861,7 +861,7 @@ \You may still want to disambiguate your grammar or your input." else (); hd results) else (ambig_msg (); error ("More than one term is type correct:\n" ^ - cat_lines (map (Pretty.string_of_term pp) (flat (map fst results))))) + cat_lines (map (Pretty.string_of_term pp) (List.concat (map fst results))))) end; @@ -1014,7 +1014,7 @@ fun class_of_const c_class = let - val c = implode (take (size c_class - size "_class", explode c_class)); + val c = implode (Library.take (size c_class - size "_class", explode c_class)); in if const_of_class c = c_class then c else raise TERM ("class_of_const: bad name " ^ quote c_class, []) @@ -1088,9 +1088,9 @@ val path' = if elems = "//" then NONE else if elems = "/" then SOME [] - else if elems = ".." andalso is_some path andalso path <> SOME [] then - SOME (fst (split_last (the path))) - else SOME (if_none path [] @ NameSpace.unpack elems); + else if elems = ".." andalso isSome path andalso path <> SOME [] then + SOME (fst (split_last (valOf path))) + else SOME (getOpt (path,[]) @ NameSpace.unpack elems); in (syn, tsig, ctab, (path', spaces), data) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/sorts.ML --- a/src/Pure/sorts.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/sorts.ML Thu Mar 03 12:43:01 2005 +0100 @@ -133,7 +133,7 @@ fun norm_sort _ [] = [] | norm_sort _ (S as [_]) = S | norm_sort classes S = - sort_strings (distinct_class (filter (minimal_class classes S) S)); + sort_strings (distinct_class (List.filter (minimal_class classes S) S)); end; @@ -161,7 +161,7 @@ in intr S end; (*instersect sorts (preserves minimality)*) -fun inter_sort classes = sort_strings o foldr (inter_class classes); +fun inter_sort classes = sort_strings o Library.foldr (inter_class classes); @@ -179,7 +179,7 @@ NONE => raise DOMAIN (a, c) | SOME Ss => Ss); val doms = map mg_dom S; - in foldl (ListPair.map (inter_sort classes)) (hd doms, tl doms) end; + in Library.foldl (ListPair.map (inter_sort classes)) (hd doms, tl doms) end; (* of_sort *) @@ -231,8 +231,8 @@ witn_types path ts (solved_failed, S) else let val ((solved', failed'), ws) = witn_sorts (S :: path) (solved_failed, SS) in - if forall is_some ws then - let val w = (Type (t, map (#1 o the) ws), S) + if forall isSome ws then + let val w = (Type (t, map (#1 o valOf) ws), S) in ((w :: solved', failed'), SOME w) end else witn_types path ts ((solved', failed'), S) end @@ -252,7 +252,7 @@ | check_result (SOME (T, S)) = if of_sort (classes, arities) (T, S) then SOME (T, S) else sys_error ("Sorts.witness_sorts: bad witness for sort " ^ str_of_sort S); - in mapfilter check_result (#2 (witness_aux (classes, arities) log_types hyps sorts)) end; + in List.mapPartial check_result (#2 (witness_aux (classes, arities) log_types hyps sorts)) end; end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/tactic.ML --- a/src/Pure/tactic.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/tactic.ML Thu Mar 03 12:43:01 2005 +0100 @@ -368,7 +368,7 @@ (*As above, but inserts only facts (unconditional theorems); generates no additional subgoals. *) -fun cut_facts_tac ths = cut_rules_tac (filter is_fact ths); +fun cut_facts_tac ths = cut_rules_tac (List.filter is_fact ths); (*Introduce the given proposition as a lemma and subgoal*) fun subgoal_tac sprop = @@ -426,7 +426,7 @@ (*build a pair of nets for biresolution*) fun build_netpair netpair brls = - foldr insert_tagged_brl (taglist 1 brls, netpair); + Library.foldr insert_tagged_brl (taglist 1 brls, netpair); (*delete one kbrl from the pair of nets*) local @@ -473,7 +473,7 @@ (*build a net of rules for resolution*) fun build_net rls = - foldr insert_krl (taglist 1 rls, Net.empty); + Library.foldr insert_krl (taglist 1 rls, Net.empty); (*resolution using a net rather than rules; pred supports filt_resolve_tac*) fun filt_resolution_from_net_tac match pred net = @@ -562,7 +562,7 @@ val rev_defs = sort_lhs_depths o map symmetric; -fun fold_rule defs thm = foldl (fn (th, ds) => rewrite_rule ds th) (thm, rev_defs defs); +fun fold_rule defs thm = Library.foldl (fn (th, ds) => rewrite_rule ds th) (thm, rev_defs defs); fun fold_tac defs = EVERY (map rewrite_tac (rev_defs defs)); fun fold_goals_tac defs = EVERY (map rewrite_goals_tac (rev_defs defs)); @@ -620,7 +620,7 @@ else (Then tac (rotate_tac n THEN' etac thin_rl),0); in SUBGOAL(fn (subg,n) => let val Hs = Logic.strip_assums_hyp subg - in case fst(foldl thins ((NONE,0),Hs)) of + in case fst(Library.foldl thins ((NONE,0),Hs)) of NONE => no_tac | SOME tac => tac n end) end; @@ -644,8 +644,8 @@ val statement = Logic.list_implies (asms, prop); val frees = map Term.dest_Free (Term.term_frees statement); val fixed_frees = filter_out (fn (x, _) => x mem_string xs) frees; - val fixed_tfrees = foldr Term.add_typ_tfree_names (map #2 fixed_frees, []); - val params = mapfilter (fn x => apsome (pair x) (assoc_string (frees, x))) xs; + val fixed_tfrees = Library.foldr Term.add_typ_tfree_names (map #2 fixed_frees, []); + val params = List.mapPartial (fn x => Option.map (pair x) (assoc_string (frees, x))) xs; fun err msg = raise ERROR_MESSAGE (msg ^ "\nThe error(s) above occurred for the goal statement:\n" ^ diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/tctical.ML --- a/src/Pure/tctical.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/tctical.ML Thu Mar 03 12:43:01 2005 +0100 @@ -179,10 +179,10 @@ fun EVERY1 tacs = EVERY' tacs 1; (* FIRST [tac1,...,tacn] equals tac1 ORELSE ... ORELSE tacn *) -fun FIRST tacs = foldr (op ORELSE) (tacs, no_tac); +fun FIRST tacs = Library.foldr (op ORELSE) (tacs, no_tac); (* FIRST' [tac1,...,tacn] i equals tac1 i ORELSE ... ORELSE tacn i *) -fun FIRST' tacs = foldr (op ORELSE') (tacs, K no_tac); +fun FIRST' tacs = Library.foldr (op ORELSE') (tacs, K no_tac); (*Apply first tactic to 1*) fun FIRST1 tacs = FIRST' tacs 1; @@ -439,7 +439,7 @@ let val {sign,maxidx,...} = rep_thm state val cterm = cterm_of sign (*find all vars in the hyps -- should find tvars also!*) - val hyps_vars = foldr add_term_vars (Logic.strip_assums_hyp prem, []) + val hyps_vars = Library.foldr add_term_vars (Logic.strip_assums_hyp prem, []) val insts = map mk_inst hyps_vars (*replace the hyps_vars by Frees*) val prem' = subst_atomic insts prem @@ -472,7 +472,7 @@ fun relift st = let val prop = #prop(rep_thm st) val subgoal_vars = (*Vars introduced in the subgoals*) - foldr add_term_vars (Logic.strip_imp_prems prop, []) + Library.foldr add_term_vars (Logic.strip_imp_prems prop, []) and concl_vars = add_term_vars (Logic.strip_imp_concl prop, []) val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st @@ -503,7 +503,7 @@ fun PRIMITIVE thmfun = PRIMSEQ (Seq.single o thmfun); (* Inverse (more or less) of PRIMITIVE *) -fun SINGLE tacf = apsome fst o Seq.pull o tacf +fun SINGLE tacf = Option.map fst o Seq.pull o tacf end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/term.ML --- a/src/Pure/term.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/term.ML Thu Mar 03 12:43:01 2005 +0100 @@ -254,7 +254,7 @@ fun S --> T = Type("fun",[S,T]); (*handy for multiple args: [T1,...,Tn]--->T gives T1-->(T2--> ... -->T)*) -val op ---> = foldr (op -->); +val op ---> = Library.foldr (op -->); fun dest_Type (Type x) = x | dest_Type T = raise TYPE ("dest_Type", [T], []); @@ -309,8 +309,8 @@ Ts = [T0,T1,...] holds types of bound variables 0, 1, ...*) fun type_of1 (Ts, Const (_,T)) = T | type_of1 (Ts, Free (_,T)) = T - | type_of1 (Ts, Bound i) = (nth_elem (i,Ts) - handle LIST _ => raise TYPE("type_of: bound variable", [], [Bound i])) + | type_of1 (Ts, Bound i) = (List.nth (Ts,i) + handle Subscript => raise TYPE("type_of: bound variable", [], [Bound i])) | type_of1 (Ts, Var (_,T)) = T | type_of1 (Ts, Abs (_,T,body)) = T --> type_of1(T::Ts, body) | type_of1 (Ts, f$u) = @@ -334,15 +334,15 @@ | _ => raise TERM("fastype_of: expected function type", [f$u])) | fastype_of1 (_, Const (_,T)) = T | fastype_of1 (_, Free (_,T)) = T - | fastype_of1 (Ts, Bound i) = (nth_elem(i,Ts) - handle LIST _ => raise TERM("fastype_of: Bound", [Bound i])) + | fastype_of1 (Ts, Bound i) = (List.nth(Ts,i) + handle Subscript => raise TERM("fastype_of: Bound", [Bound i])) | fastype_of1 (_, Var (_,T)) = T | fastype_of1 (Ts, Abs (_,T,u)) = T --> fastype_of1 (T::Ts, u); fun fastype_of t : typ = fastype_of1 ([],t); -val list_abs = foldr (fn ((x, T), t) => Abs (x, T, t)); +val list_abs = Library.foldr (fn ((x, T), t) => Abs (x, T, t)); (* maps (x1,...,xn)t to t *) fun strip_abs_body (Abs(_,_,t)) = strip_abs_body t @@ -365,7 +365,7 @@ (* maps (f, [t1,...,tn]) to f(t1,...,tn) *) -val list_comb : term * term list -> term = foldl (op $); +val list_comb : term * term list -> term = Library.foldl (op $); (* maps f(t1,...,tn) to (f, [t1,...,tn]) ; naturally tail-recursive*) @@ -465,7 +465,7 @@ let val ren = match_bvs (pat, obj, []); fun ren_abs (Abs (x, T, b)) = - Abs (if_none (assoc_string (ren, x)) x, T, ren_abs b) + Abs (getOpt (assoc_string (ren, x), x), T, ren_abs b) | ren_abs (f $ t) = ren_abs f $ ren_abs t | ren_abs t = t in if null ren then NONE else SOME (ren_abs t) end; @@ -645,14 +645,14 @@ | subst_atomic (instl: (term*term) list) t = let fun subst (Abs(a,T,body)) = Abs(a, T, subst body) | subst (f$t') = subst f $ subst t' - | subst t = if_none (assoc(instl,t)) t + | subst t = getOpt (assoc(instl,t),t) in subst t end; (*Substitute for type Vars in a type*) fun typ_subst_TVars iTs T = if null iTs then T else let fun subst(Type(a,Ts)) = Type(a, map subst Ts) | subst(T as TFree _) = T - | subst(T as TVar(ixn,_)) = if_none (assoc(iTs,ixn)) T + | subst(T as TVar(ixn,_)) = getOpt (assoc(iTs,ixn),T) in subst T end; (*Substitute for type Vars in a term*) @@ -660,7 +660,7 @@ (*Substitute for Vars in a term; see also envir/norm_term*) fun subst_Vars itms t = if null itms then t else - let fun subst(v as Var(ixn,_)) = if_none (assoc(itms,ixn)) v + let fun subst(v as Var(ixn,_)) = getOpt (assoc(itms,ixn),v) | subst(Abs(a,T,t)) = Abs(a,T,subst t) | subst(f$t) = subst f $ subst t | subst(t) = t @@ -695,7 +695,7 @@ | maxidx_of_term (Abs (_,T,u)) = Int.max(maxidx_of_term u, maxidx_of_typ T) | maxidx_of_term (f$t) = Int.max(maxidx_of_term f, maxidx_of_term t); -fun maxidx_of_terms ts = foldl Int.max (~1, map maxidx_of_term ts); +fun maxidx_of_terms ts = Library.foldl Int.max (~1, map maxidx_of_term ts); (* Increment the index of all Poly's in T by k *) @@ -712,7 +712,7 @@ let fun check (T as Type ("dummy", _)) = raise TYPE ("Illegal occurrence of '_' dummy type", [T], []) - | check (Type (_, Ts)) = seq check Ts + | check (Type (_, Ts)) = List.app check Ts | check _ = (); in check typ; typ end; @@ -747,11 +747,11 @@ (** Consts etc. **) -fun add_typ_classes (Type (_, Ts), cs) = foldr add_typ_classes (Ts, cs) +fun add_typ_classes (Type (_, Ts), cs) = Library.foldr add_typ_classes (Ts, cs) | add_typ_classes (TFree (_, S), cs) = S union cs | add_typ_classes (TVar (_, S), cs) = S union cs; -fun add_typ_tycons (Type (c, Ts), cs) = foldr add_typ_tycons (Ts, c ins cs) +fun add_typ_tycons (Type (c, Ts), cs) = Library.foldr add_typ_tycons (Ts, c ins cs) | add_typ_tycons (_, cs) = cs; val add_term_classes = it_term_types add_typ_classes; @@ -814,20 +814,20 @@ | add_term_names (_, bs) = bs; (*Accumulates the TVars in a type, suppressing duplicates. *) -fun add_typ_tvars(Type(_,Ts),vs) = foldr add_typ_tvars (Ts,vs) +fun add_typ_tvars(Type(_,Ts),vs) = Library.foldr add_typ_tvars (Ts,vs) | add_typ_tvars(TFree(_),vs) = vs | add_typ_tvars(TVar(v),vs) = v ins vs; (*Accumulates the TFrees in a type, suppressing duplicates. *) -fun add_typ_tfree_names(Type(_,Ts),fs) = foldr add_typ_tfree_names (Ts,fs) +fun add_typ_tfree_names(Type(_,Ts),fs) = Library.foldr add_typ_tfree_names (Ts,fs) | add_typ_tfree_names(TFree(f,_),fs) = f ins_string fs | add_typ_tfree_names(TVar(_),fs) = fs; -fun add_typ_tfrees(Type(_,Ts),fs) = foldr add_typ_tfrees (Ts,fs) +fun add_typ_tfrees(Type(_,Ts),fs) = Library.foldr add_typ_tfrees (Ts,fs) | add_typ_tfrees(TFree(f),fs) = f ins fs | add_typ_tfrees(TVar(_),fs) = fs; -fun add_typ_varnames(Type(_,Ts),nms) = foldr add_typ_varnames (Ts,nms) +fun add_typ_varnames(Type(_,Ts),nms) = Library.foldr add_typ_varnames (Ts,nms) | add_typ_varnames(TFree(nm,_),nms) = nm ins_string nms | add_typ_varnames(TVar((nm,_),_),nms) = nm ins_string nms; @@ -848,7 +848,7 @@ (*special code to enforce left-to-right collection of TVar-indexnames*) -fun add_typ_ixns(ixns,Type(_,Ts)) = foldl add_typ_ixns (ixns,Ts) +fun add_typ_ixns(ixns,Type(_,Ts)) = Library.foldl add_typ_ixns (ixns,Ts) | add_typ_ixns(ixns,TVar(ixn,_)) = if mem_ix (ixn, ixns) then ixns else ixns@[ixn] | add_typ_ixns(ixns,TFree(_)) = ixns; @@ -908,7 +908,7 @@ fun rename_aTs names vars : (string*typ)list = let fun rename_aT (vars,(a,T)) = (variant (map #1 vars @ names) a, T) :: vars - in foldl rename_aT ([],vars) end; + in Library.foldl rename_aT ([],vars) end; fun rename_wrt_term t = rename_aTs (add_term_names(t,[])); @@ -916,7 +916,7 @@ (* left-ro-right traversal *) (*foldl atoms of type*) -fun foldl_atyps f (x, Type (_, Ts)) = foldl (foldl_atyps f) (x, Ts) +fun foldl_atyps f (x, Type (_, Ts)) = Library.foldl (foldl_atyps f) (x, Ts) | foldl_atyps f x_atom = f x_atom; (*foldl atoms of term*) @@ -1094,7 +1094,7 @@ fun string_of_vname (x, i) = let val si = string_of_int i; - val dot = if_none (try (Symbol.is_digit o last_elem o Symbol.explode) x) true; + val dot = getOpt (try (Symbol.is_digit o List.last o Symbol.explode) x, true); val qmark = if !show_var_qmarks then "?" else ""; in if dot then qmark ^ x ^ "." ^ si diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/theory.ML --- a/src/Pure/theory.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/theory.ML Thu Mar 03 12:43:01 2005 +0100 @@ -333,7 +333,7 @@ end handle TERM _ => NONE; fun clash_defns c_ty axms = - distinct (mapfilter (clash_defn c_ty) axms); + distinct (List.mapPartial (clash_defn c_ty) axms); (* overloading *) @@ -396,7 +396,7 @@ fun add_deps (c, cs) deps = let fun declare (G, x) = Graph.new_node (x, ()) G handle Graph.DUP _ => G - in foldl declare (deps, c :: cs) |> Graph.add_deps_acyclic (c, cs) end; + in Library.foldl declare (deps, c :: cs) |> Graph.add_deps_acyclic (c, cs) end; fun check_defn sg overloaded final_consts ((deps, axms), def as (name, tm)) = let @@ -455,10 +455,10 @@ fun ext_defns prep_axm overloaded raw_axms thy = let val Theory {sign, const_deps, final_consts, axioms, oracles, parents, ancestors} = thy; - val all_axioms = flat (map (Symtab.dest o #axioms o rep_theory) (thy :: ancestors)); + val all_axioms = List.concat (map (Symtab.dest o #axioms o rep_theory) (thy :: ancestors)); val axms = map (prep_axm sign) raw_axms; - val (const_deps', _) = foldl (check_defn sign overloaded final_consts) ((const_deps, all_axioms), axms); + val (const_deps', _) = Library.foldl (check_defn sign overloaded final_consts) ((const_deps, all_axioms), axms); in make_theory sign const_deps' final_consts axioms oracles parents ancestors |> add_axioms_i axms @@ -502,7 +502,7 @@ | NONE => Symtab.update((c,[ty]),finals)) end; val consts = map (prep_term sign) raw_terms; - val final_consts' = foldl mk_final (final_consts,consts); + val final_consts' = Library.foldl mk_final (final_consts,consts); in make_theory sign const_deps final_consts' axioms oracles parents ancestors end; @@ -523,7 +523,7 @@ else (ty::gen_rem (Sign.typ_instance sg) (tys,ty)); fun merge ([],_) = [] | merge (_,[]) = [] - | merge input = foldl merge_single input; + | merge input = Library.foldl merge_single input; in SOME o merge end; @@ -552,28 +552,28 @@ else let val sign' = - foldl merge_sign (sign_of (hd thys), tl thys) + Library.foldl merge_sign (sign_of (hd thys), tl thys) |> Sign.prep_ext |> Sign.add_path "/"; val depss = map (#const_deps o rep_theory) thys; - val deps' = foldl (Graph.merge_acyclic (K true)) (hd depss, tl depss) + val deps' = Library.foldl (Graph.merge_acyclic (K true)) (hd depss, tl depss) handle Graph.CYCLES namess => error (cycle_msg namess); val final_constss = map (#final_consts o rep_theory) thys; - val final_consts' = foldl (Symtab.join (merge_final sign')) (hd final_constss, + val final_consts' = Library.foldl (Symtab.join (merge_final sign')) (hd final_constss, tl final_constss); val axioms' = Symtab.empty; fun eq_ora ((_, (_, s1: stamp)), (_, (_, s2))) = s1 = s2; val oracles' = Symtab.make (gen_distinct eq_ora - (flat (map (Symtab.dest o #oracles o rep_theory) thys))) + (List.concat (map (Symtab.dest o #oracles o rep_theory) thys))) handle Symtab.DUPS names => err_dup_oras names; val parents' = gen_distinct eq_thy thys; val ancestors' = - gen_distinct eq_thy (parents' @ flat (map ancestors_of thys)); + gen_distinct eq_thy (parents' @ List.concat (map ancestors_of thys)); in make_theory sign' deps' final_consts' axioms' oracles' parents' ancestors' end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/thm.ML --- a/src/Pure/thm.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/thm.ML Thu Mar 03 12:43:01 2005 +0100 @@ -204,7 +204,7 @@ (*Destruct abstraction in cterms*) fun dest_abs a (Cterm {sign_ref, T as Type("fun",[_,S]), maxidx, t=Abs(x,ty,M)}) = - let val (y,N) = variant_abs (if_none a x,ty,M) + let val (y,N) = variant_abs (getOpt (a,x),ty,M) in (Cterm {sign_ref = sign_ref, T = ty, maxidx = 0, t = Free(y,ty)}, Cterm {sign_ref = sign_ref, T = S, maxidx = maxidx, t = N}) end @@ -244,7 +244,7 @@ val vars = map dest_Var (term_vars t1); fun mk_cTinsts (ixn, T) = (ixn, Ctyp {sign_ref = sign_ref, T = T}); fun mk_ctinsts (ixn, t) = - let val T = typ_subst_TVars Tinsts (the (assoc (vars, ixn))) + let val T = typ_subst_TVars Tinsts (valOf (assoc (vars, ixn))) in (Cterm {sign_ref = sign_ref, maxidx = maxidx, T = T, t = Var (ixn, T)}, Cterm {sign_ref = sign_ref, maxidx = maxidx, T = T, t = t}) @@ -299,7 +299,7 @@ tpairs: (term * term) list, (*flex-flex pairs*) prop: term}; (*conclusion*) -fun terms_of_tpairs tpairs = flat (map (op @ o pairself single) tpairs); +fun terms_of_tpairs tpairs = List.concat (map (op @ o pairself single) tpairs); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); @@ -439,7 +439,7 @@ (* fix_shyps *) -fun all_sorts_nonempty sign_ref = is_some (Sign.universal_witness (Sign.deref sign_ref)); +fun all_sorts_nonempty sign_ref = isSome (Sign.universal_witness (Sign.deref sign_ref)); (*preserve sort contexts of rule premises and substituted types*) fun fix_shyps thms Ts (thm as Thm {sign_ref, der, maxidx, hyps, tpairs, prop, ...}) = @@ -922,7 +922,7 @@ let val ntpairs = map (pairself (Envir.norm_term env)) tpairs; val newprop = Envir.norm_term env prop; (*Remove trivial tpairs, of the form t=t*) - val distpairs = filter (not o op aconv) ntpairs + val distpairs = List.filter (not o op aconv) ntpairs in fix_shyps [th] (env_codT env) (Thm{sign_ref = sign_ref, der = Pt.infer_derivs' (Pt.norm_proof' env) der, @@ -977,8 +977,8 @@ No longer normalizes the new theorem! *) fun instantiate ([], []) th = th | instantiate (vcTs,ctpairs) (th as Thm{sign_ref,der,maxidx,hyps,shyps,tpairs=dpairs,prop}) = - let val (newsign_ref,tpairs) = foldr add_ctpair (ctpairs, (sign_ref,[])); - val (newsign_ref,vTs) = foldr add_ctyp (vcTs, (newsign_ref,[])); + let val (newsign_ref,tpairs) = Library.foldr add_ctpair (ctpairs, (sign_ref,[])); + val (newsign_ref,vTs) = Library.foldr add_ctyp (vcTs, (newsign_ref,[])); fun subst t = subst_atomic tpairs (Sign.inst_term_tvars (Sign.deref newsign_ref) vTs t); val newprop = subst prop; @@ -1041,7 +1041,7 @@ (* Replace all TFrees not fixed or in the hyps by new TVars *) fun varifyT' fixed (Thm{sign_ref,der,maxidx,shyps,hyps,tpairs,prop}) = let - val tfrees = foldr add_term_tfree_names (hyps, fixed); + val tfrees = Library.foldr add_term_tfree_names (hyps, fixed); val prop1 = attach_tpairs tpairs prop; val (prop2, al) = Type.varify (prop1, tfrees); val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2) @@ -1232,7 +1232,7 @@ val short = length iparams - length cs val newnames = if short<0 then error"More names than abstractions!" - else variantlist(take (short,iparams), cs) @ cs + else variantlist(Library.take (short,iparams), cs) @ cs val freenames = map (#1 o dest_Free) (term_frees Bi) val newBi = Logic.list_rename_params (newnames, Bi) in @@ -1282,7 +1282,7 @@ Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs([],_,_,_) = I | rename_bvs(al,dpairs,tpairs,B) = - let val vars = foldr add_term_vars + let val vars = Library.foldr add_term_vars (map fst dpairs @ map fst tpairs @ map snd tpairs, []) (*unknowns appearing elsewhere be preserved!*) val vids = map (#1 o #1 o dest_Var) vars; @@ -1292,7 +1292,7 @@ else Var((y,i),T) | NONE=> t) | rename(Abs(x,T,t)) = - Abs(if_none(assoc_string(al,x)) x, T, rename t) + Abs(getOpt(assoc_string(al,x),x), T, rename t) | rename(f$t) = rename f $ rename t | rename(t) = t; fun strip_ren Ai = strip_apply rename (Ai,B) @@ -1300,7 +1300,7 @@ (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars(dpairs, tpairs, B) = - rename_bvs(foldr Term.match_bvars (dpairs,[]), dpairs, tpairs, B); + rename_bvs(Library.foldr Term.match_bvars (dpairs,[]), dpairs, tpairs, B); (*** RESOLUTION ***) diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/type.ML --- a/src/Pure/type.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/type.ML Thu Mar 03 12:43:01 2005 +0100 @@ -147,7 +147,7 @@ local fun inst_typ env (Type (c, Ts)) = Type (c, map (inst_typ env) Ts) - | inst_typ env (T as TFree (x, _)) = if_none (Library.assoc_string (env, x)) T + | inst_typ env (T as TFree (x, _)) = getOpt (Library.assoc_string (env, x),T) | inst_typ _ T = T; fun certify_typ normalize syntax tsig ty = @@ -238,11 +238,11 @@ in ((ix,v)::pairs, v::used) end; fun freeze_one alist (ix,sort) = - TFree (the (assoc (alist, ix)), sort) + TFree (valOf (assoc (alist, ix)), sort) handle Option => raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []); -fun thaw_one alist (a,sort) = TVar (the (assoc (alist,a)), sort) +fun thaw_one alist (a,sort) = TVar (valOf (assoc (alist,a)), sort) handle Option => TFree(a, sort); in @@ -252,14 +252,14 @@ let val used = add_typ_tfree_names (T, []) and tvars = map #1 (add_typ_tvars (T, [])); - val (alist, _) = foldr new_name (tvars, ([], used)); + val (alist, _) = Library.foldr new_name (tvars, ([], used)); in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end; fun freeze_thaw t = let val used = it_term_types add_typ_tfree_names (t, []) and tvars = map #1 (it_term_types add_typ_tvars (t, [])); - val (alist, _) = foldr new_name (tvars, ([], used)); + val (alist, _) = Library.foldr new_name (tvars, ([], used)); in (case alist of [] => (t, fn x => x) (*nothing to do!*) @@ -303,7 +303,7 @@ | SOME U => if U = T then subs else raise TYPE_MATCH) | match (subs, (Type (a, Ts), Type (b, Us))) = if a <> b then raise TYPE_MATCH - else foldl match (subs, Ts ~~ Us) + else Library.foldl match (subs, Ts ~~ Us) | match (subs, (TFree x, TFree y)) = if x = y then subs else raise TYPE_MATCH | match _ = raise TYPE_MATCH; @@ -378,7 +378,7 @@ else meet ((T, S), Vartab.update_new ((v, T), tye)) | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY - else foldr unif (Ts ~~ Us, tye) + else Library.foldr unif (Ts ~~ Us, tye) | (T, U) => if T = U then tye else raise TUNIFY); in (unif (TU, tyenv), ! tyvar_count) end; @@ -434,7 +434,7 @@ fun insert_arities pp classes (arities, (t, ars)) = let val ars' = Symtab.lookup_multi (arities, t) - |> curry (foldr (insert pp classes t)) (flat (map (complete classes) ars)) + |> curry (Library.foldr (insert pp classes t)) (List.concat (map (complete classes) ars)) in Symtab.update ((t, ars'), arities) end; fun insert_table pp classes = Symtab.foldl (fn (arities, (t, ars)) => @@ -455,7 +455,7 @@ | NONE => error (undecl_type t)); val ars = decls |> map ((fn (t, Ss, S) => (t, map (fn c => (c, Ss)) S)) o prep); - val arities' = foldl (insert_arities pp classes) (arities, ars); + val arities' = Library.foldl (insert_arities pp classes) (arities, ars); in (classes, default, types, arities') end); fun rebuild_arities pp classes arities = @@ -533,7 +533,7 @@ SOME (decl', _) => err_in_decls c decl decl' | NONE => Symtab.update ((c, (decl, stamp ())), types)); -fun the_decl types c = fst (the (Symtab.lookup (types, c))); +fun the_decl types c = fst (valOf (Symtab.lookup (types, c))); fun change_types f = change_tsig (fn (classes, default, types, arities) => (classes, default, f types, arities)); diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/type_infer.ML --- a/src/Pure/type_infer.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/type_infer.ML Thu Mar 03 12:43:01 2005 +0100 @@ -114,7 +114,7 @@ if is_param xi andalso is_none (assoc (ps, xi)) then (xi, mk_param S) :: ps else ps | add_parms (ps, TFree _) = ps - | add_parms (ps, Type (_, Ts)) = foldl add_parms (ps, Ts); + | add_parms (ps, Type (_, Ts)) = Library.foldl add_parms (ps, Ts); val params' = add_parms (params, typ); @@ -149,7 +149,7 @@ | add_vparms (ps, _) = ps; val vparams' = add_vparms (vparams, tm); - fun var_param xi = the (assoc (vparams', xi)); + fun var_param xi = valOf (assoc (vparams', xi)); val preT_of = pretyp_of is_param; @@ -192,7 +192,7 @@ (* add_parms *) -fun add_parmsT (rs, PType (_, Ts)) = foldl add_parmsT (rs, Ts) +fun add_parmsT (rs, PType (_, Ts)) = Library.foldl add_parmsT (rs, Ts) | add_parmsT (rs, Link (r as ref (Param _))) = r ins rs | add_parmsT (rs, Link (ref T)) = add_parmsT (rs, T) | add_parmsT (rs, _) = rs; @@ -202,7 +202,7 @@ (* add_names *) -fun add_namesT (xs, PType (_, Ts)) = foldl add_namesT (xs, Ts) +fun add_namesT (xs, PType (_, Ts)) = Library.foldl add_namesT (xs, Ts) | add_namesT (xs, PTFree (x, _)) = x ins xs | add_namesT (xs, PTVar ((x, _), _)) = x ins xs | add_namesT (xs, Link (ref T)) = add_namesT (xs, T) @@ -237,8 +237,8 @@ fun elim (r as ref (Param S), x) = r := mk_var (x, S) | elim _ = (); - val used' = foldl add_names (foldl add_namesT (used, Ts), ts); - val parms = rev (foldl add_parms (foldl add_parmsT ([], Ts), ts)); + val used' = Library.foldl add_names (Library.foldl add_namesT (used, Ts), ts); + val parms = rev (Library.foldl add_parms (Library.foldl add_parmsT ([], Ts), ts)); val names = Term.invent_names used' (prfx ^ "'a") (length parms); in seq2 elim (parms, names); @@ -285,7 +285,7 @@ fun occurs_check r (Link (r' as ref T)) = if r = r' then raise NO_UNIFIER "Occurs check!" else occurs_check r T - | occurs_check r (PType (_, Ts)) = seq (occurs_check r) Ts + | occurs_check r (PType (_, Ts)) = List.app (occurs_check r) Ts | occurs_check _ _ = (); fun assign r T S = @@ -376,7 +376,7 @@ fun inf _ (PConst (_, T)) = T | inf _ (PFree (_, T)) = T | inf _ (PVar (_, T)) = T - | inf bs (PBound i) = snd (nth_elem (i, bs) handle LIST _ => err_loose i) + | inf bs (PBound i) = snd (List.nth (bs, i) handle Subscript => err_loose i) | inf bs (PAbs (x, T, t)) = PType ("fun", [T, inf ((x, T) :: bs) t]) | inf bs (PAppl (t, u)) = let @@ -405,12 +405,12 @@ (*run type inference*) val tTs' = ListPair.map Constraint (ts', Ts'); - val _ = seq (fn t => (infer pp classes arities t; ())) tTs'; + val _ = List.app (fn t => (infer pp classes arities t; ())) tTs'; (*collect result unifier*) fun ch_var (xi, Link (r as ref (Param S))) = (r := PTVar (xi, S); NONE) | ch_var xi_T = SOME xi_T; - val env = mapfilter ch_var Tps; + val env = List.mapPartial ch_var Tps; (*convert back to terms/typs*) val mk_var = @@ -472,8 +472,8 @@ fun decode_types tsig is_const def_type def_sort map_const map_type map_sort tm = let - fun get_type xi = if_none (def_type xi) dummyT; - fun is_free x = is_some (def_type (x, ~1)); + fun get_type xi = getOpt (def_type xi, dummyT); + fun is_free x = isSome (def_type (x, ~1)); val raw_env = Syntax.raw_term_sorts tm; val sort_of = get_sort tsig def_sort map_sort raw_env; @@ -519,7 +519,7 @@ let val {classes, arities, ...} = Type.rep_tsig tsig; val pat_Ts' = map (Type.cert_typ tsig) pat_Ts; - val is_const = is_some o const_type; + val is_const = isSome o const_type; val raw_ts' = map (decode_types tsig is_const def_type def_sort map_const map_type map_sort) raw_ts; val (ts, Ts, unifier) = basic_infer_types pp const_type diff -r 1b3115d1a8df -r 8d8c70b41bab src/Pure/unify.ML --- a/src/Pure/unify.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/Pure/unify.ML Thu Mar 03 12:43:01 2005 +0100 @@ -291,7 +291,7 @@ Clever would be to re-do just the affected dpairs*) fun SIMPL (env,dpairs) : Envir.env * dpair list * dpair list = let val all as (env',flexflex,flexrigid) = - foldr SIMPL0 (dpairs, (env,[],[])); + Library.foldr SIMPL0 (dpairs, (env,[],[])); val dps = flexrigid@flexflex in if exists (fn ((_,t,u)) => changed(env',t) orelse changed(env',u)) dps then SIMPL(env',dps) else all @@ -458,7 +458,7 @@ if i j < i-lev) banned)) + else Bound (i - length (List.filter (fn j => j < i-lev) banned)) | change lev (Abs (a,T,t)) = Abs (a, T, change(lev+1) t) | change lev (t$u) = change lev t $ change lev u | change lev t = t @@ -488,7 +488,7 @@ val (Ts,U) = strip_type env T and js = length ts - 1 downto 0 val args = sort (make_ord arg_less) - (foldr (change_arg banned) (flexargs (js,ts,Ts), [])) + (Library.foldr (change_arg banned) (flexargs (js,ts,Ts), [])) val ts' = map (#t) args in if decreasing (length Ts) args then (env, (list_comb(Var(v,T), ts'))) @@ -521,7 +521,7 @@ then (bnos, (a,T)::newbinder) (*needed by both: keep*) else (j::bnos, newbinder); (*remove*) val indices = 0 upto (length rbinder - 1); - val (banned,rbin') = foldr add_index (rbinder~~indices, ([],[])); + val (banned,rbin') = Library.foldr add_index (rbinder~~indices, ([],[])); val (env', t') = clean_term banned (env, t); val (env'',u') = clean_term banned (env',u) in (ff_assign(env'', rbin', t', u'), tpairs) @@ -560,7 +560,7 @@ val bsymbs = [termT u, Pretty.str" =?=", Pretty.brk 1, termT t]; in tracing(Pretty.string_of(Pretty.blk(0,bsymbs))) end; - in tracing msg; seq pdp dpairs end; + in tracing msg; List.app pdp dpairs end; (*Unify the dpairs in the environment. @@ -575,7 +575,7 @@ then print_dpairs "Enter SIMPL" (env,dpairs) else (); SIMPL (env,dpairs)) in case flexrigid of - [] => SOME (foldr add_ffpair (flexflex, (env',[])), reseq) + [] => SOME (Library.foldr add_ffpair (flexflex, (env',[])), reseq) | dp::frigid' => if tdepth > !search_bound then (warning "Unification bound exceeded"; Seq.pull reseq) @@ -626,7 +626,7 @@ (*Smash all flex-flexpairs. Should allow selection of pairs by a predicate?*) fun smash_flexflex (env,tpairs) : Envir.env = - foldr smash_flexflex1 (tpairs, env); + Library.foldr smash_flexflex1 (tpairs, env); (*Returns unifiers with no remaining disagreement pairs*) fun smash_unifiers (sg, env, tus) : Envir.env Seq.seq = diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Datatype.ML --- a/src/ZF/Datatype.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Datatype.ML Thu Mar 03 12:43:01 2005 +0100 @@ -73,9 +73,9 @@ and (rhead, rargs) = strip_comb rhs val lname = #1 (dest_Const lhead) handle TERM _ => raise Match; val rname = #1 (dest_Const rhead) handle TERM _ => raise Match; - val lcon_info = the (Symtab.lookup (ConstructorsData.get_sg sg, lname)) + val lcon_info = valOf (Symtab.lookup (ConstructorsData.get_sg sg, lname)) handle Option => raise Match; - val rcon_info = the (Symtab.lookup (ConstructorsData.get_sg sg, rname)) + val rcon_info = valOf (Symtab.lookup (ConstructorsData.get_sg sg, rname)) handle Option => raise Match; val new = if #big_rec_name lcon_info = #big_rec_name rcon_info diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Tools/cartprod.ML --- a/src/ZF/Tools/cartprod.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Tools/cartprod.ML Thu Mar 03 12:43:01 2005 +0100 @@ -96,7 +96,7 @@ (*Makes a nested tuple from a list, following the product type structure*) fun mk_tuple pair (Type("*", [T1,T2])) tms = pair $ (mk_tuple pair T1 tms) - $ (mk_tuple pair T2 (drop (length (factors T1), tms))) + $ (mk_tuple pair T2 (Library.drop (length (factors T1), tms))) | mk_tuple pair T (t::_) = t; (*Attempts to remove occurrences of split, and pair-valued parameters*) diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Tools/datatype_package.ML --- a/src/ZF/Tools/datatype_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Tools/datatype_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -126,20 +126,20 @@ (*Treatment of a list of constructors, for one part Result adds a list of terms, each a function variable with arguments*) fun add_case_list (con_ty_list, (opno, case_lists)) = - let val (opno', case_list) = foldr add_case (con_ty_list, (opno, [])) + let val (opno', case_list) = Library.foldr add_case (con_ty_list, (opno, [])) in (opno', case_list :: case_lists) end; (*Treatment of all parts*) - val (_, case_lists) = foldr add_case_list (con_ty_lists, (1,[])); + val (_, case_lists) = Library.foldr add_case_list (con_ty_lists, (1,[])); (*extract the types of all the variables*) - val case_typ = flat (map (map (#2 o #1)) con_ty_lists) ---> (iT-->iT); + val case_typ = List.concat (map (map (#2 o #1)) con_ty_lists) ---> (iT-->iT); val case_base_name = big_rec_base_name ^ "_case"; val case_name = full_name case_base_name; (*The list of all the function variables*) - val case_args = flat (map (map #1) case_lists); + val case_args = List.concat (map (map #1) case_lists); val case_const = Const (case_name, case_typ); val case_tm = list_comb (case_const, case_args); @@ -170,7 +170,7 @@ val rec_args = map (make_rec_call (rev case_args,0)) (List.drop(recursor_args, ncase_args)) in - foldr add_abs + Library.foldr add_abs (case_args, list_comb (recursor_var, bound_args @ rec_args)) end @@ -202,23 +202,23 @@ val rec_ty_lists = (map (map rec_ty_elem) con_ty_lists); (*Treatment of all parts*) - val (_, recursor_lists) = foldr add_case_list (rec_ty_lists, (1,[])); + val (_, recursor_lists) = Library.foldr add_case_list (rec_ty_lists, (1,[])); (*extract the types of all the variables*) - val recursor_typ = flat (map (map (#2 o #1)) rec_ty_lists) + val recursor_typ = List.concat (map (map (#2 o #1)) rec_ty_lists) ---> (iT-->iT); val recursor_base_name = big_rec_base_name ^ "_rec"; val recursor_name = full_name recursor_base_name; (*The list of all the function variables*) - val recursor_args = flat (map (map #1) recursor_lists); + val recursor_args = List.concat (map (map #1) recursor_lists); val recursor_tm = list_comb (Const (recursor_name, recursor_typ), recursor_args); val recursor_cases = map call_recursor - (flat case_lists ~~ flat recursor_lists) + (List.concat case_lists ~~ List.concat recursor_lists) val recursor_def = Logic.mk_defpair @@ -240,16 +240,16 @@ val (thy0, con_defs) = thy_path |> Theory.add_consts_i ((case_base_name, case_typ, NoSyn) :: - map #1 (flat con_ty_lists)) + map #1 (List.concat con_ty_lists)) |> PureThy.add_defs_i false (map Thm.no_attributes (case_def :: - flat (ListPair.map mk_con_defs + List.concat (ListPair.map mk_con_defs (1 upto npart, con_ty_lists)))) |>> add_recursor |>> Theory.parent_path - val intr_names = map #2 (flat con_ty_lists); + val intr_names = map #2 (List.concat con_ty_lists); val (thy1, ind_result) = thy0 |> Ind_Package.add_inductive_i false (rec_tms, dom_sum) (map Thm.no_attributes (intr_names ~~ intr_tms)) @@ -291,7 +291,7 @@ val case_eqns = map prove_case_eqn - (flat con_ty_lists ~~ case_args ~~ tl con_defs); + (List.concat con_ty_lists ~~ case_args ~~ tl con_defs); (*** Prove the recursor theorems ***) @@ -316,7 +316,7 @@ (recursor_tm $ (list_comb (Const (Sign.intern_const (sign_of thy1) name,T), args)), - subst_rec (foldl betapply (recursor_case, args)))); + subst_rec (Library.foldl betapply (recursor_case, args)))); val recursor_trans = recursor_def RS def_Vrecursor RS trans; @@ -332,7 +332,7 @@ (cterm_of (sign_of thy1) (mk_recursor_eqn arg))) recursor_tacsf in - map prove_recursor_eqn (flat con_ty_lists ~~ recursor_cases) + map prove_recursor_eqn (List.concat con_ty_lists ~~ recursor_cases) end val constructors = @@ -384,7 +384,7 @@ (("free_iffs", free_iffs), []), (("free_elims", free_SEs), [])]) |> DatatypesData.map (fn tab => Symtab.update ((big_rec_name, dt_info), tab)) - |> ConstructorsData.map (fn tab => foldr Symtab.update (con_pairs, tab)) + |> ConstructorsData.map (fn tab => Library.foldr Symtab.update (con_pairs, tab)) |> Theory.parent_path, ind_result, {con_defs = con_defs, diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Tools/induct_tacs.ML --- a/src/ZF/Tools/induct_tacs.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Tools/induct_tacs.ML Thu Mar 03 12:43:01 2005 +0100 @@ -89,7 +89,7 @@ let fun mk_pair (Const("op :",_) $ Free (v,_) $ A) = (v, #1 (dest_Const (head_of A))) | mk_pair _ = raise Match - val pairs = mapfilter (try (mk_pair o FOLogic.dest_Trueprop)) + val pairs = List.mapPartial (try (mk_pair o FOLogic.dest_Trueprop)) (#2 (strip_context Bi)) in case assoc (pairs, var) of NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var) @@ -178,7 +178,7 @@ (Symtab.update ((big_rec_name, dt_info), DatatypesData.get thy)) |> ConstructorsData.put - (foldr Symtab.update (con_pairs, ConstructorsData.get thy)) + (Library.foldr Symtab.update (con_pairs, ConstructorsData.get thy)) |> Theory.parent_path end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Tools/inductive_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -102,7 +102,7 @@ Sign.string_of_term sign t); (*** Construct the fixedpoint definition ***) - val mk_variant = variant (foldr add_term_names (intr_tms, [])); + val mk_variant = variant (Library.foldr add_term_names (intr_tms, [])); val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w"; @@ -113,10 +113,10 @@ (*Makes a disjunct from an introduction rule*) fun fp_part intr = (*quantify over rule's free vars except parameters*) let val prems = map dest_tprop (strip_imp_prems intr) - val dummy = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds + val dummy = List.app (fn rec_hd => List.app (chk_prem rec_hd) prems) rec_hds val exfrees = term_frees intr \\ rec_params val zeq = FOLogic.mk_eq (Free(z',iT), #1 (rule_concl intr)) - in foldr FOLogic.mk_exists + in Library.foldr FOLogic.mk_exists (exfrees, fold_bal FOLogic.mk_conj (zeq::prems)) end; @@ -138,7 +138,7 @@ val fp_rhs = Fp.oper $ dom_sum $ fp_abs - val dummy = seq (fn rec_hd => deny (rec_hd occs fp_rhs) + val dummy = List.app (fn rec_hd => deny (rec_hd occs fp_rhs) "Illegal occurrence of recursion operator") rec_hds; @@ -173,7 +173,7 @@ (*tracing: print the fixedpoint definition*) val dummy = if !Ind_Syntax.trace then - seq (writeln o Sign.string_of_term sign o #2) axpairs + List.app (writeln o Sign.string_of_term sign o #2) axpairs else () (*add definitions of the inductive sets*) @@ -311,7 +311,7 @@ (*Make a premise of the induction rule.*) fun induct_prem ind_alist intr = let val quantfrees = map dest_Free (term_frees intr \\ rec_params) - val iprems = foldr (add_induct_prem ind_alist) + val iprems = Library.foldr (add_induct_prem ind_alist) (Logic.strip_imp_prems intr,[]) val (t,X) = Ind_Syntax.rule_concl intr val (SOME pred) = gen_assoc (op aconv) (ind_alist, X) @@ -332,7 +332,7 @@ val dummy = if !Ind_Syntax.trace then (writeln "ind_prems = "; - seq (writeln o Sign.string_of_term sign1) ind_prems; + List.app (writeln o Sign.string_of_term sign1) ind_prems; writeln "raw_induct = "; print_thm raw_induct) else (); @@ -390,7 +390,7 @@ val pfree = Free(pred_name ^ "_" ^ Sign.base_name rec_name, elem_factors ---> FOLogic.oT) val qconcl = - foldr FOLogic.mk_all + Library.foldr FOLogic.mk_all (elem_frees, FOLogic.imp $ (Ind_Syntax.mem_const $ elem_tuple $ rec_tm) diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Tools/primrec_package.ML --- a/src/ZF/Tools/primrec_package.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Tools/primrec_package.ML Thu Mar 03 12:43:01 2005 +0100 @@ -56,7 +56,7 @@ else strip_comb (hd middle); val (cname, _) = dest_Const constr handle TERM _ => raise RecError "ill-formed constructor"; - val con_info = the (Symtab.lookup (ConstructorsData.get thy, cname)) + val con_info = valOf (Symtab.lookup (ConstructorsData.get thy, cname)) handle Option => raise RecError "cannot determine datatype associated with function" @@ -80,7 +80,7 @@ else case rec_fn_opt of NONE => SOME (fname, ftype, ls, rs, con_info, [new_eqn]) | SOME (fname', _, ls', rs', con_info': constructor_info, eqns) => - if is_some (assoc (eqns, cname)) then + if isSome (assoc (eqns, cname)) then raise RecError "constructor already occurred as pattern" else if (ls <> ls') orelse (rs <> rs') then raise RecError "non-recursive arguments are inconsistent" @@ -128,7 +128,7 @@ | SOME (rhs, cargs', eq) => (rhs, inst_recursor (recursor_pair, cargs'), eq) val allowed_terms = map use_fabs (#2 (strip_comb recursor_rhs)) - val abs = foldr absterm (allowed_terms, rhs) + val abs = Library.foldr absterm (allowed_terms, rhs) in if !Ind_Syntax.trace then writeln ("recursor_rhs = " ^ @@ -153,7 +153,7 @@ val def_tm = Logic.mk_equals (subst_bound (rec_arg, fabs), list_comb (recursor, - foldr add_case (cnames ~~ recursor_pairs, [])) + Library.foldr add_case (cnames ~~ recursor_pairs, [])) $ rec_arg) in @@ -172,7 +172,7 @@ let val ((eqn_names, eqn_terms), eqn_atts) = apfst split_list (split_list args); val SOME (fname, ftype, ls, rs, con_info, eqns) = - foldr (process_eqn thy) (eqn_terms, NONE); + Library.foldr (process_eqn thy) (eqn_terms, NONE); val def = process_fun thy (fname, ftype, ls, rs, con_info, eqns); val (thy1, [def_thm]) = thy diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/Tools/typechk.ML --- a/src/ZF/Tools/typechk.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/Tools/typechk.ML Thu Mar 03 12:43:01 2005 +0100 @@ -68,8 +68,8 @@ else (warning ("No such type-checking rule\n" ^ (string_of_thm th)); cs); -val op addTCs = foldl addTC; -val op delTCs = foldl delTC; +val op addTCs = Library.foldl addTC; +val op delTCs = Library.foldl delTC; (*resolution using a net rather than rules*) diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/arith_data.ML --- a/src/ZF/arith_data.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/arith_data.ML Thu Mar 03 12:43:01 2005 +0100 @@ -71,7 +71,7 @@ fun prove_conv name tacs sg hyps xs (t,u) = if t aconv u then NONE else - let val hyps' = filter (not o is_eq_thm) hyps + let val hyps' = List.filter (not o is_eq_thm) hyps val goal = Logic.list_implies (map (#prop o Thm.rep_thm) hyps', FOLogic.mk_Trueprop (mk_eq_iff (t, u))); in SOME (hyps' MRS Tactic.prove sg xs [] goal (K (EVERY tacs))) diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/ind_syntax.ML --- a/src/ZF/ind_syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/ind_syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -111,7 +111,7 @@ fun union_params (rec_tm, cs) = let val (_,args) = strip_comb rec_tm fun is_ind arg = (type_of arg = iT) - in case filter is_ind (args @ cs) of + in case List.filter is_ind (args @ cs) of [] => empty | u_args => fold_bal mk_Un u_args end; @@ -134,8 +134,8 @@ foldl_aterms (fn (cs, Const c) => Const c ins cs | (cs, _) => cs); fun fourth (_, name, args, prems) = prems fun add_consts_in_cts cts = - foldl (foldl add_term_consts_2) ([], map fourth (flat cts)); - val cs = filter is_OK (add_consts_in_cts con_ty_lists) + Library.foldl (Library.foldl add_term_consts_2) ([], map fourth (List.concat cts)); + val cs = List.filter is_OK (add_consts_in_cts con_ty_lists) in u $ union_params (hd rec_tms, cs) end; diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/simpdata.ML --- a/src/ZF/simpdata.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/simpdata.ML Thu Mar 03 12:43:01 2005 +0100 @@ -16,7 +16,7 @@ case head_of t of Const(a,_) => (case assoc(pairs,a) of - SOME rls => flat (map (atomize (conn_pairs, mem_pairs)) + SOME rls => List.concat (map (atomize (conn_pairs, mem_pairs)) ([th] RL rls)) | NONE => [th]) | _ => [th] diff -r 1b3115d1a8df -r 8d8c70b41bab src/ZF/thy_syntax.ML --- a/src/ZF/thy_syntax.ML Thu Mar 03 09:22:35 2005 +0100 +++ b/src/ZF/thy_syntax.ML Thu Mar 03 12:43:01 2005 +0100 @@ -90,7 +90,7 @@ val big_rec_name = space_implode "_" rec_names and srec_tms = mk_list (map fst rec_pairs) and scons = mk_scons rec_pairs - val con_names = flat (map (map (unenclose o #1 o #1) o snd) + val con_names = List.concat (map (map (unenclose o #1 o #1) o snd) rec_pairs) val inames = mk_list (map (mk_bind "_I") con_names) in