--- a/src/Tools/Metis/src/Active.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-(* ========================================================================= *)
-(* THE ACTIVE SET OF CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Active =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type of active clause sets. *)
-(* ------------------------------------------------------------------------- *)
-
-type simplify = {subsume : bool, reduce : bool, rewrite : bool}
-
-type parameters =
- {clause : Clause.parameters,
- prefactor : simplify,
- postfactor : simplify}
-
-type active
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val default : parameters
-
-val size : active -> int
-
-val saturated : active -> Clause.clause list
-
-(* ------------------------------------------------------------------------- *)
-(* Create a new active clause set and initialize clauses. *)
-(* ------------------------------------------------------------------------- *)
-
-val new : parameters -> Thm.thm list -> active * Clause.clause list
-
-(* ------------------------------------------------------------------------- *)
-(* Add a clause into the active set and deduce all consequences. *)
-(* ------------------------------------------------------------------------- *)
-
-val add : active -> Clause.clause -> active * Clause.clause list
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val pp : active Parser.pp
-
-end
--- a/src/Tools/Metis/src/Active.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,803 +0,0 @@
-(* ========================================================================= *)
-(* THE ACTIVE SET OF CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Active :> Active =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun allFactors red =
- let
- fun allClause cl = List.all red (cl :: Clause.factor cl)
- in
- List.all allClause
- end;
-
- fun allResolutions red =
- let
- fun allClause2 cl_lit cl =
- let
- fun allLiteral2 lit =
- case total (Clause.resolve cl_lit) (cl,lit) of
- NONE => true
- | SOME cl => allFactors red [cl]
- in
- LiteralSet.all allLiteral2 (Clause.literals cl)
- end
-
- fun allClause1 allCls cl =
- let
- val cl = Clause.freshVars cl
-
- fun allLiteral1 lit = List.all (allClause2 (cl,lit)) allCls
- in
- LiteralSet.all allLiteral1 (Clause.literals cl)
- end
- in
- fn [] => true
- | allCls as cl :: cls =>
- allClause1 allCls cl andalso allResolutions red cls
- end;
-
- fun allParamodulations red cls =
- let
- fun allClause2 cl_lit_ort_tm cl =
- let
- fun allLiteral2 lit =
- let
- val para = Clause.paramodulate cl_lit_ort_tm
-
- fun allSubterms (path,tm) =
- case total para (cl,lit,path,tm) of
- NONE => true
- | SOME cl => allFactors red [cl]
- in
- List.all allSubterms (Literal.nonVarTypedSubterms lit)
- end
- in
- LiteralSet.all allLiteral2 (Clause.literals cl)
- end
-
- fun allClause1 cl =
- let
- val cl = Clause.freshVars cl
-
- fun allLiteral1 lit =
- let
- fun allCl2 x = List.all (allClause2 x) cls
- in
- case total Literal.destEq lit of
- NONE => true
- | SOME (l,r) =>
- allCl2 (cl,lit,Rewrite.LeftToRight,l) andalso
- allCl2 (cl,lit,Rewrite.RightToLeft,r)
- end
- in
- LiteralSet.all allLiteral1 (Clause.literals cl)
- end
- in
- List.all allClause1 cls
- end;
-
- fun redundant {subsume,reduce,rewrite} =
- let
- fun simp cl =
- case Clause.simplify cl of
- NONE => true
- | SOME cl =>
- Subsume.isStrictlySubsumed subsume (Clause.literals cl) orelse
- let
- val cl' = cl
- val cl' = Clause.reduce reduce cl'
- val cl' = Clause.rewrite rewrite cl'
- in
- not (Clause.equalThms cl cl') andalso simp cl'
- end
- in
- simp
- end;
-in
- fun isSaturated ordering subs cls =
- let
-(*TRACE2
- val ppCls = Parser.ppList Clause.pp
- val () = Parser.ppTrace ppCls "Active.checkSaturated: clauses" cls
-*)
- val red = Units.empty
- val rw = Rewrite.new (KnuthBendixOrder.compare ordering)
- val red = redundant {subsume = subs, reduce = red, rewrite = rw}
- in
- allFactors red cls andalso
- allResolutions red cls andalso
- allParamodulations red cls
- end;
-
- fun checkSaturated ordering subs cls =
- if isSaturated ordering subs cls then () else raise Bug "unsaturated";
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of active clause sets. *)
-(* ------------------------------------------------------------------------- *)
-
-type simplify = {subsume : bool, reduce : bool, rewrite : bool};
-
-type parameters =
- {clause : Clause.parameters,
- prefactor : simplify,
- postfactor : simplify};
-
-datatype active =
- Active of
- {parameters : parameters,
- clauses : Clause.clause IntMap.map,
- units : Units.units,
- rewrite : Rewrite.rewrite,
- subsume : Clause.clause Subsume.subsume,
- literals : (Clause.clause * Literal.literal) LiteralNet.literalNet,
- equations :
- (Clause.clause * Literal.literal * Rewrite.orient * Term.term)
- TermNet.termNet,
- subterms :
- (Clause.clause * Literal.literal * Term.path * Term.term)
- TermNet.termNet,
- allSubterms : (Clause.clause * Term.term) TermNet.termNet};
-
-fun getSubsume (Active {subsume = s, ...}) = s;
-
-fun setRewrite active rewrite =
- let
- val Active
- {parameters,clauses,units,subsume,literals,equations,
- subterms,allSubterms,...} = active
- in
- Active
- {parameters = parameters, clauses = clauses, units = units,
- rewrite = rewrite, subsume = subsume, literals = literals,
- equations = equations, subterms = subterms, allSubterms = allSubterms}
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val maxSimplify : simplify = {subsume = true, reduce = true, rewrite = true};
-
-val default : parameters =
- {clause = Clause.default,
- prefactor = maxSimplify,
- postfactor = maxSimplify};
-
-fun empty parameters =
- let
- val {clause,...} = parameters
- val {ordering,...} = clause
- in
- Active
- {parameters = parameters,
- clauses = IntMap.new (),
- units = Units.empty,
- rewrite = Rewrite.new (KnuthBendixOrder.compare ordering),
- subsume = Subsume.new (),
- literals = LiteralNet.new {fifo = false},
- equations = TermNet.new {fifo = false},
- subterms = TermNet.new {fifo = false},
- allSubterms = TermNet.new {fifo = false}}
- end;
-
-fun size (Active {clauses,...}) = IntMap.size clauses;
-
-fun clauses (Active {clauses = cls, ...}) =
- let
- fun add (_,cl,acc) = cl :: acc
- in
- IntMap.foldr add [] cls
- end;
-
-fun saturated active =
- let
- fun remove (cl,(cls,subs)) =
- let
- val lits = Clause.literals cl
- in
- if Subsume.isStrictlySubsumed subs lits then (cls,subs)
- else (cl :: cls, Subsume.insert subs (lits,()))
- end
-
- val cls = clauses active
- val (cls,_) = foldl remove ([], Subsume.new ()) cls
- val (cls,subs) = foldl remove ([], Subsume.new ()) cls
-
-(*DEBUG
- val Active {parameters,...} = active
- val {clause,...} = parameters
- val {ordering,...} = clause
- val () = checkSaturated ordering subs cls
-*)
- in
- cls
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val pp =
- let
- fun toStr active = "Active{" ^ Int.toString (size active) ^ "}"
- in
- Parser.ppMap toStr Parser.ppString
- end;
-
-(*DEBUG
-local
- open Parser;
-
- fun ppField f ppA p a =
- (beginBlock p Inconsistent 2;
- addString p (f ^ " =");
- addBreak p (1,0);
- ppA p a;
- endBlock p);
-
- val ppClauses =
- ppField "clauses"
- (Parser.ppMap IntMap.toList
- (Parser.ppList (Parser.ppPair Parser.ppInt Clause.pp)));
-
- val ppRewrite = ppField "rewrite" Rewrite.pp;
-
- val ppSubterms =
- ppField "subterms"
- (TermNet.pp
- (Parser.ppMap (fn (c,l,p,t) => ((Clause.id c, l, p), t))
- (Parser.ppPair
- (Parser.ppTriple Parser.ppInt Literal.pp Term.ppPath)
- Term.pp)));
-in
- fun pp p (Active {clauses,rewrite,subterms,...}) =
- (beginBlock p Inconsistent 2;
- addString p "Active";
- addBreak p (1,0);
- beginBlock p Inconsistent 1;
- addString p "{";
- ppClauses p clauses;
- addString p ",";
- addBreak p (1,0);
- ppRewrite p rewrite;
-(*TRACE5
- addString p ",";
- addBreak p (1,0);
- ppSubterms p subterms;
-*)
- endBlock p;
- addString p "}";
- endBlock p);
-end;
-*)
-
-val toString = Parser.toString pp;
-
-(* ------------------------------------------------------------------------- *)
-(* Simplify clauses. *)
-(* ------------------------------------------------------------------------- *)
-
-fun simplify simp units rewr subs =
- let
- val {subsume = s, reduce = r, rewrite = w} = simp
-
- fun rewrite cl =
- let
- val cl' = Clause.rewrite rewr cl
- in
- if Clause.equalThms cl cl' then SOME cl else Clause.simplify cl'
- end
- in
- fn cl =>
- case Clause.simplify cl of
- NONE => NONE
- | SOME cl =>
- case (if w then rewrite cl else SOME cl) of
- NONE => NONE
- | SOME cl =>
- let
- val cl = if r then Clause.reduce units cl else cl
- in
- if s andalso Clause.subsumes subs cl then NONE else SOME cl
- end
- end;
-
-(*DEBUG
-val simplify = fn simp => fn units => fn rewr => fn subs => fn cl =>
- let
- fun traceCl s = Parser.ppTrace Clause.pp ("Active.simplify: " ^ s)
-(*TRACE4
- val ppClOpt = Parser.ppOption Clause.pp
- val () = traceCl "cl" cl
-*)
- val cl' = simplify simp units rewr subs cl
-(*TRACE4
- val () = Parser.ppTrace ppClOpt "Active.simplify: cl'" cl'
-*)
- val () =
- case cl' of
- NONE => ()
- | SOME cl' =>
- case
- (case simplify simp units rewr subs cl' of
- NONE => SOME ("away", K ())
- | SOME cl'' =>
- if Clause.equalThms cl' cl'' then NONE
- else SOME ("further", fn () => traceCl "cl''" cl'')) of
- NONE => ()
- | SOME (e,f) =>
- let
- val () = traceCl "cl" cl
- val () = traceCl "cl'" cl'
- val () = f ()
- in
- raise
- Bug
- ("Active.simplify: clause should have been simplified "^e)
- end
- in
- cl'
- end;
-*)
-
-fun simplifyActive simp active =
- let
- val Active {units,rewrite,subsume,...} = active
- in
- simplify simp units rewrite subsume
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Add a clause into the active set. *)
-(* ------------------------------------------------------------------------- *)
-
-fun addUnit units cl =
- let
- val th = Clause.thm cl
- in
- case total Thm.destUnit th of
- SOME lit => Units.add units (lit,th)
- | NONE => units
- end;
-
-fun addRewrite rewrite cl =
- let
- val th = Clause.thm cl
- in
- case total Thm.destUnitEq th of
- SOME l_r => Rewrite.add rewrite (Clause.id cl, (l_r,th))
- | NONE => rewrite
- end;
-
-fun addSubsume subsume cl = Subsume.insert subsume (Clause.literals cl, cl);
-
-fun addLiterals literals cl =
- let
- fun add (lit as (_,atm), literals) =
- if Atom.isEq atm then literals
- else LiteralNet.insert literals (lit,(cl,lit))
- in
- LiteralSet.foldl add literals (Clause.largestLiterals cl)
- end;
-
-fun addEquations equations cl =
- let
- fun add ((lit,ort,tm),equations) =
- TermNet.insert equations (tm,(cl,lit,ort,tm))
- in
- foldl add equations (Clause.largestEquations cl)
- end;
-
-fun addSubterms subterms cl =
- let
- fun add ((lit,path,tm),subterms) =
- TermNet.insert subterms (tm,(cl,lit,path,tm))
- in
- foldl add subterms (Clause.largestSubterms cl)
- end;
-
-fun addAllSubterms allSubterms cl =
- let
- fun add ((_,_,tm),allSubterms) =
- TermNet.insert allSubterms (tm,(cl,tm))
- in
- foldl add allSubterms (Clause.allSubterms cl)
- end;
-
-fun addClause active cl =
- let
- val Active
- {parameters,clauses,units,rewrite,subsume,literals,
- equations,subterms,allSubterms} = active
- val clauses = IntMap.insert clauses (Clause.id cl, cl)
- and subsume = addSubsume subsume cl
- and literals = addLiterals literals cl
- and equations = addEquations equations cl
- and subterms = addSubterms subterms cl
- and allSubterms = addAllSubterms allSubterms cl
- in
- Active
- {parameters = parameters, clauses = clauses, units = units,
- rewrite = rewrite, subsume = subsume, literals = literals,
- equations = equations, subterms = subterms,
- allSubterms = allSubterms}
- end;
-
-fun addFactorClause active cl =
- let
- val Active
- {parameters,clauses,units,rewrite,subsume,literals,
- equations,subterms,allSubterms} = active
- val units = addUnit units cl
- and rewrite = addRewrite rewrite cl
- in
- Active
- {parameters = parameters, clauses = clauses, units = units,
- rewrite = rewrite, subsume = subsume, literals = literals,
- equations = equations, subterms = subterms, allSubterms = allSubterms}
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Derive (unfactored) consequences of a clause. *)
-(* ------------------------------------------------------------------------- *)
-
-fun deduceResolution literals cl (lit as (_,atm), acc) =
- let
- fun resolve (cl_lit,acc) =
- case total (Clause.resolve cl_lit) (cl,lit) of
- SOME cl' => cl' :: acc
- | NONE => acc
-(*TRACE4
- val () = Parser.ppTrace Literal.pp "Active.deduceResolution: lit" lit
-*)
- in
- if Atom.isEq atm then acc
- else foldl resolve acc (LiteralNet.unify literals (Literal.negate lit))
- end;
-
-fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) =
- let
- fun para (cl_lit_path_tm,acc) =
- case total (Clause.paramodulate (cl,lit,ort,tm)) cl_lit_path_tm of
- SOME cl' => cl' :: acc
- | NONE => acc
- in
- foldl para acc (TermNet.unify subterms tm)
- end;
-
-fun deduceParamodulationInto equations cl ((lit,path,tm),acc) =
- let
- fun para (cl_lit_ort_tm,acc) =
- case total (Clause.paramodulate cl_lit_ort_tm) (cl,lit,path,tm) of
- SOME cl' => cl' :: acc
- | NONE => acc
- in
- foldl para acc (TermNet.unify equations tm)
- end;
-
-fun deduce active cl =
- let
- val Active {parameters,literals,equations,subterms,...} = active
-
- val lits = Clause.largestLiterals cl
- val eqns = Clause.largestEquations cl
- val subtms =
- if TermNet.null equations then [] else Clause.largestSubterms cl
-
- val acc = []
- val acc = LiteralSet.foldl (deduceResolution literals cl) acc lits
- val acc = foldl (deduceParamodulationWith subterms cl) acc eqns
- val acc = foldl (deduceParamodulationInto equations cl) acc subtms
- in
- rev acc
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Extract clauses from the active set that can be simplified. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun clause_rewritables active =
- let
- val Active {clauses,rewrite,...} = active
-
- fun rewr (id,cl,ids) =
- let
- val cl' = Clause.rewrite rewrite cl
- in
- if Clause.equalThms cl cl' then ids else IntSet.add ids id
- end
- in
- IntMap.foldr rewr IntSet.empty clauses
- end;
-
- fun orderedRedexResidues (((l,r),_),ort) =
- case ort of
- NONE => []
- | SOME Rewrite.LeftToRight => [(l,r,true)]
- | SOME Rewrite.RightToLeft => [(r,l,true)];
-
- fun unorderedRedexResidues (((l,r),_),ort) =
- case ort of
- NONE => [(l,r,false),(r,l,false)]
- | SOME _ => [];
-
- fun rewrite_rewritables active rewr_ids =
- let
- val Active {parameters,rewrite,clauses,allSubterms,...} = active
- val {clause = {ordering,...}, ...} = parameters
- val order = KnuthBendixOrder.compare ordering
-
- fun addRewr (id,acc) =
- if IntMap.inDomain id clauses then IntSet.add acc id else acc
-
- fun addReduce ((l,r,ord),acc) =
- let
- fun isValidRewr tm =
- case total (Subst.match Subst.empty l) tm of
- NONE => false
- | SOME sub =>
- ord orelse
- let
- val tm' = Subst.subst (Subst.normalize sub) r
- in
- order (tm,tm') = SOME GREATER
- end
-
- fun addRed ((cl,tm),acc) =
- let
-(*TRACE5
- val () = Parser.ppTrace Clause.pp "Active.addRed: cl" cl
- val () = Parser.ppTrace Term.pp "Active.addRed: tm" tm
-*)
- val id = Clause.id cl
- in
- if IntSet.member id acc then acc
- else if not (isValidRewr tm) then acc
- else IntSet.add acc id
- end
-
-(*TRACE5
- val () = Parser.ppTrace Term.pp "Active.addReduce: l" l
- val () = Parser.ppTrace Term.pp "Active.addReduce: r" r
- val () = Parser.ppTrace Parser.ppBool "Active.addReduce: ord" ord
-*)
- in
- List.foldl addRed acc (TermNet.matched allSubterms l)
- end
-
- fun addEquation redexResidues (id,acc) =
- case Rewrite.peek rewrite id of
- NONE => acc
- | SOME eqn_ort => List.foldl addReduce acc (redexResidues eqn_ort)
-
- val addOrdered = addEquation orderedRedexResidues
-
- val addUnordered = addEquation unorderedRedexResidues
-
- val ids = IntSet.empty
- val ids = List.foldl addRewr ids rewr_ids
- val ids = List.foldl addOrdered ids rewr_ids
- val ids = List.foldl addUnordered ids rewr_ids
- in
- ids
- end;
-
- fun choose_clause_rewritables active ids = size active <= length ids
-
- fun rewritables active ids =
- if choose_clause_rewritables active ids then clause_rewritables active
- else rewrite_rewritables active ids;
-
-(*DEBUG
- val rewritables = fn active => fn ids =>
- let
- val clause_ids = clause_rewritables active
- val rewrite_ids = rewrite_rewritables active ids
-
- val () =
- if IntSet.equal rewrite_ids clause_ids then ()
- else
- let
- val ppIdl = Parser.ppList Parser.ppInt
- val ppIds = Parser.ppMap IntSet.toList ppIdl
- val () = Parser.ppTrace pp "Active.rewritables: active" active
- val () = Parser.ppTrace ppIdl "Active.rewritables: ids" ids
- val () = Parser.ppTrace ppIds
- "Active.rewritables: clause_ids" clause_ids
- val () = Parser.ppTrace ppIds
- "Active.rewritables: rewrite_ids" rewrite_ids
- in
- raise Bug "Active.rewritables: ~(rewrite_ids SUBSET clause_ids)"
- end
- in
- if choose_clause_rewritables active ids then clause_ids else rewrite_ids
- end;
-*)
-
- fun delete active ids =
- if IntSet.null ids then active
- else
- let
- fun idPred id = not (IntSet.member id ids)
-
- fun clausePred cl = idPred (Clause.id cl)
-
- val Active
- {parameters,clauses,units,rewrite,subsume,literals,
- equations,subterms,allSubterms} = active
-
- val clauses = IntMap.filter (idPred o fst) clauses
- and subsume = Subsume.filter clausePred subsume
- and literals = LiteralNet.filter (clausePred o #1) literals
- and equations = TermNet.filter (clausePred o #1) equations
- and subterms = TermNet.filter (clausePred o #1) subterms
- and allSubterms = TermNet.filter (clausePred o fst) allSubterms
- in
- Active
- {parameters = parameters, clauses = clauses, units = units,
- rewrite = rewrite, subsume = subsume, literals = literals,
- equations = equations, subterms = subterms,
- allSubterms = allSubterms}
- end;
-in
- fun extract_rewritables (active as Active {clauses,rewrite,...}) =
- if Rewrite.isReduced rewrite then (active,[])
- else
- let
-(*TRACE3
- val () = trace "Active.extract_rewritables: inter-reducing\n"
-*)
- val (rewrite,ids) = Rewrite.reduce' rewrite
- val active = setRewrite active rewrite
- val ids = rewritables active ids
- val cls = IntSet.transform (IntMap.get clauses) ids
-(*TRACE3
- val ppCls = Parser.ppList Clause.pp
- val () = Parser.ppTrace ppCls "Active.extract_rewritables: cls" cls
-*)
- in
- (delete active ids, cls)
- end
-(*DEBUG
- handle Error err =>
- raise Bug ("Active.extract_rewritables: shouldn't fail\n" ^ err);
-*)
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Factor clauses. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun prefactor_simplify active subsume =
- let
- val Active {parameters,units,rewrite,...} = active
- val {prefactor,...} = parameters
- in
- simplify prefactor units rewrite subsume
- end;
-
- fun postfactor_simplify active subsume =
- let
- val Active {parameters,units,rewrite,...} = active
- val {postfactor,...} = parameters
- in
- simplify postfactor units rewrite subsume
- end;
-
- val sort_utilitywise =
- let
- fun utility cl =
- case LiteralSet.size (Clause.literals cl) of
- 0 => ~1
- | 1 => if Thm.isUnitEq (Clause.thm cl) then 0 else 1
- | n => n
- in
- sortMap utility Int.compare
- end;
-
- fun factor_add (cl, active_subsume_acc as (active,subsume,acc)) =
- case postfactor_simplify active subsume cl of
- NONE => active_subsume_acc
- | SOME cl =>
- let
- val active = addFactorClause active cl
- and subsume = addSubsume subsume cl
- and acc = cl :: acc
- in
- (active,subsume,acc)
- end;
-
- fun factor1 (cl, active_subsume_acc as (active,subsume,_)) =
- case prefactor_simplify active subsume cl of
- NONE => active_subsume_acc
- | SOME cl =>
- let
- val cls = sort_utilitywise (cl :: Clause.factor cl)
- in
- foldl factor_add active_subsume_acc cls
- end;
-
- fun factor' active acc [] = (active, rev acc)
- | factor' active acc cls =
- let
- val cls = sort_utilitywise cls
- val subsume = getSubsume active
- val (active,_,acc) = foldl factor1 (active,subsume,acc) cls
- val (active,cls) = extract_rewritables active
- in
- factor' active acc cls
- end;
-in
- fun factor active cls = factor' active [] cls;
-end;
-
-(*TRACE4
-val factor = fn active => fn cls =>
- let
- val ppCls = Parser.ppList Clause.pp
- val () = Parser.ppTrace ppCls "Active.factor: cls" cls
- val (active,cls') = factor active cls
- val () = Parser.ppTrace ppCls "Active.factor: cls'" cls'
- in
- (active,cls')
- end;
-*)
-
-(* ------------------------------------------------------------------------- *)
-(* Create a new active clause set and initialize clauses. *)
-(* ------------------------------------------------------------------------- *)
-
-fun new parameters ths =
- let
- val {clause,...} = parameters
-
- fun mk_clause th =
- Clause.mk {parameters = clause, id = Clause.newId (), thm = th}
-
- val cls = map mk_clause ths
- in
- factor (empty parameters) cls
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Add a clause into the active set and deduce all consequences. *)
-(* ------------------------------------------------------------------------- *)
-
-fun add active cl =
- case simplifyActive maxSimplify active cl of
- NONE => (active,[])
- | SOME cl' =>
- if Clause.isContradiction cl' then (active,[cl'])
- else if not (Clause.equalThms cl cl') then factor active [cl']
- else
- let
-(*TRACE3
- val () = Parser.ppTrace Clause.pp "Active.add: cl" cl
-*)
- val active = addClause active cl
- val cl = Clause.freshVars cl
- val cls = deduce active cl
- val (active,cls) = factor active cls
-(*TRACE2
- val ppCls = Parser.ppList Clause.pp
- val () = Parser.ppTrace ppCls "Active.add: cls" cls
-*)
- in
- (active,cls)
- end;
-
-end
--- a/src/Tools/Metis/src/Atom.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,137 +0,0 @@
-(* ========================================================================= *)
-(* FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Atom =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type for storing first order logic atoms. *)
-(* ------------------------------------------------------------------------- *)
-
-type relationName = Name.name
-
-type relation = relationName * int
-
-type atom = relationName * Term.term list
-
-(* ------------------------------------------------------------------------- *)
-(* Constructors and destructors. *)
-(* ------------------------------------------------------------------------- *)
-
-val name : atom -> relationName
-
-val arguments : atom -> Term.term list
-
-val arity : atom -> int
-
-val relation : atom -> relation
-
-val functions : atom -> NameAritySet.set
-
-val functionNames : atom -> NameSet.set
-
-(* Binary relations *)
-
-val mkBinop : relationName -> Term.term * Term.term -> atom
-
-val destBinop : relationName -> atom -> Term.term * Term.term
-
-val isBinop : relationName -> atom -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* The size of an atom in symbols. *)
-(* ------------------------------------------------------------------------- *)
-
-val symbols : atom -> int
-
-(* ------------------------------------------------------------------------- *)
-(* A total comparison function for atoms. *)
-(* ------------------------------------------------------------------------- *)
-
-val compare : atom * atom -> order
-
-(* ------------------------------------------------------------------------- *)
-(* Subterms. *)
-(* ------------------------------------------------------------------------- *)
-
-val subterm : atom -> Term.path -> Term.term
-
-val subterms : atom -> (Term.path * Term.term) list
-
-val replace : atom -> Term.path * Term.term -> atom
-
-val find : (Term.term -> bool) -> atom -> Term.path option
-
-(* ------------------------------------------------------------------------- *)
-(* Free variables. *)
-(* ------------------------------------------------------------------------- *)
-
-val freeIn : Term.var -> atom -> bool
-
-val freeVars : atom -> NameSet.set
-
-(* ------------------------------------------------------------------------- *)
-(* Substitutions. *)
-(* ------------------------------------------------------------------------- *)
-
-val subst : Subst.subst -> atom -> atom
-
-(* ------------------------------------------------------------------------- *)
-(* Matching. *)
-(* ------------------------------------------------------------------------- *)
-
-val match : Subst.subst -> atom -> atom -> Subst.subst (* raises Error *)
-
-(* ------------------------------------------------------------------------- *)
-(* Unification. *)
-(* ------------------------------------------------------------------------- *)
-
-val unify : Subst.subst -> atom -> atom -> Subst.subst (* raises Error *)
-
-(* ------------------------------------------------------------------------- *)
-(* The equality relation. *)
-(* ------------------------------------------------------------------------- *)
-
-val eqRelation : relation
-
-val mkEq : Term.term * Term.term -> atom
-
-val destEq : atom -> Term.term * Term.term
-
-val isEq : atom -> bool
-
-val mkRefl : Term.term -> atom
-
-val destRefl : atom -> Term.term
-
-val isRefl : atom -> bool
-
-val sym : atom -> atom (* raises Error if given a refl *)
-
-val lhs : atom -> Term.term
-
-val rhs : atom -> Term.term
-
-(* ------------------------------------------------------------------------- *)
-(* Special support for terms with type annotations. *)
-(* ------------------------------------------------------------------------- *)
-
-val typedSymbols : atom -> int
-
-val nonVarTypedSubterms : atom -> (Term.path * Term.term) list
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val pp : atom Parser.pp
-
-val toString : atom -> string
-
-val fromString : string -> atom
-
-val parse : Term.term Parser.quotation -> atom
-
-end
--- a/src/Tools/Metis/src/Atom.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,245 +0,0 @@
-(* ========================================================================= *)
-(* FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Atom :> Atom =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* A type for storing first order logic atoms. *)
-(* ------------------------------------------------------------------------- *)
-
-type relationName = Name.name;
-
-type relation = relationName * int;
-
-type atom = relationName * Term.term list;
-
-(* ------------------------------------------------------------------------- *)
-(* Constructors and destructors. *)
-(* ------------------------------------------------------------------------- *)
-
-fun name ((rel,_) : atom) = rel;
-
-fun arguments ((_,args) : atom) = args;
-
-fun arity atm = length (arguments atm);
-
-fun relation atm = (name atm, arity atm);
-
-val functions =
- let
- fun f (tm,acc) = NameAritySet.union (Term.functions tm) acc
- in
- fn atm => foldl f NameAritySet.empty (arguments atm)
- end;
-
-val functionNames =
- let
- fun f (tm,acc) = NameSet.union (Term.functionNames tm) acc
- in
- fn atm => foldl f NameSet.empty (arguments atm)
- end;
-
-(* Binary relations *)
-
-fun mkBinop p (a,b) : atom = (p,[a,b]);
-
-fun destBinop p (x,[a,b]) =
- if x = p then (a,b) else raise Error "Atom.destBinop: wrong binop"
- | destBinop _ _ = raise Error "Atom.destBinop: not a binop";
-
-fun isBinop p = can (destBinop p);
-
-(* ------------------------------------------------------------------------- *)
-(* The size of an atom in symbols. *)
-(* ------------------------------------------------------------------------- *)
-
-fun symbols atm = foldl (fn (tm,z) => Term.symbols tm + z) 1 (arguments atm);
-
-(* ------------------------------------------------------------------------- *)
-(* A total comparison function for atoms. *)
-(* ------------------------------------------------------------------------- *)
-
-fun compare ((p1,tms1),(p2,tms2)) =
- case Name.compare (p1,p2) of
- LESS => LESS
- | EQUAL => lexCompare Term.compare (tms1,tms2)
- | GREATER => GREATER;
-
-(* ------------------------------------------------------------------------- *)
-(* Subterms. *)
-(* ------------------------------------------------------------------------- *)
-
-fun subterm _ [] = raise Bug "Atom.subterm: empty path"
- | subterm ((_,tms) : atom) (h :: t) =
- if h >= length tms then raise Error "Atom.subterm: bad path"
- else Term.subterm (List.nth (tms,h)) t;
-
-fun subterms ((_,tms) : atom) =
- let
- fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Term.subterms tm) @ l
- in
- foldl f [] (enumerate tms)
- end;
-
-fun replace _ ([],_) = raise Bug "Atom.replace: empty path"
- | replace (atm as (rel,tms)) (h :: t, res) : atom =
- if h >= length tms then raise Error "Atom.replace: bad path"
- else
- let
- val tm = List.nth (tms,h)
- val tm' = Term.replace tm (t,res)
- in
- if Sharing.pointerEqual (tm,tm') then atm
- else (rel, updateNth (h,tm') tms)
- end;
-
-fun find pred =
- let
- fun f (i,tm) =
- case Term.find pred tm of
- SOME path => SOME (i :: path)
- | NONE => NONE
- in
- fn (_,tms) : atom => first f (enumerate tms)
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Free variables. *)
-(* ------------------------------------------------------------------------- *)
-
-fun freeIn v atm = List.exists (Term.freeIn v) (arguments atm);
-
-val freeVars =
- let
- fun f (tm,acc) = NameSet.union (Term.freeVars tm) acc
- in
- fn atm => foldl f NameSet.empty (arguments atm)
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Substitutions. *)
-(* ------------------------------------------------------------------------- *)
-
-fun subst sub (atm as (p,tms)) : atom =
- let
- val tms' = Sharing.map (Subst.subst sub) tms
- in
- if Sharing.pointerEqual (tms',tms) then atm else (p,tms')
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Matching. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun matchArg ((tm1,tm2),sub) = Subst.match sub tm1 tm2;
-in
- fun match sub (p1,tms1) (p2,tms2) =
- let
- val _ = (p1 = p2 andalso length tms1 = length tms2) orelse
- raise Error "Atom.match"
- in
- foldl matchArg sub (zip tms1 tms2)
- end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Unification. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun unifyArg ((tm1,tm2),sub) = Subst.unify sub tm1 tm2;
-in
- fun unify sub (p1,tms1) (p2,tms2) =
- let
- val _ = (p1 = p2 andalso length tms1 = length tms2) orelse
- raise Error "Atom.unify"
- in
- foldl unifyArg sub (zip tms1 tms2)
- end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* The equality relation. *)
-(* ------------------------------------------------------------------------- *)
-
-val eqName = "=";
-
-val eqArity = 2;
-
-val eqRelation = (eqName,eqArity);
-
-val mkEq = mkBinop eqName;
-
-fun destEq x = destBinop eqName x;
-
-fun isEq x = isBinop eqName x;
-
-fun mkRefl tm = mkEq (tm,tm);
-
-fun destRefl atm =
- let
- val (l,r) = destEq atm
- val _ = l = r orelse raise Error "Atom.destRefl"
- in
- l
- end;
-
-fun isRefl x = can destRefl x;
-
-fun sym atm =
- let
- val (l,r) = destEq atm
- val _ = l <> r orelse raise Error "Atom.sym: refl"
- in
- mkEq (r,l)
- end;
-
-fun lhs atm = fst (destEq atm);
-
-fun rhs atm = snd (destEq atm);
-
-(* ------------------------------------------------------------------------- *)
-(* Special support for terms with type annotations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun typedSymbols ((_,tms) : atom) =
- foldl (fn (tm,z) => Term.typedSymbols tm + z) 1 tms;
-
-fun nonVarTypedSubterms (_,tms) =
- let
- fun addArg ((n,arg),acc) =
- let
- fun addTm ((path,tm),acc) = (n :: path, tm) :: acc
- in
- foldl addTm acc (Term.nonVarTypedSubterms arg)
- end
- in
- foldl addArg [] (enumerate tms)
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val pp = Parser.ppMap Term.Fn Term.pp;
-
-val toString = Parser.toString pp;
-
-fun fromString s = Term.destFn (Term.fromString s);
-
-val parse = Parser.parseQuotation Term.toString fromString;
-
-end
-
-structure AtomOrdered =
-struct type t = Atom.atom val compare = Atom.compare end
-
-structure AtomSet = ElementSet (AtomOrdered);
-
-structure AtomMap = KeyMap (AtomOrdered);
--- a/src/Tools/Metis/src/AtomNet.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,48 +0,0 @@
-(* ========================================================================= *)
-(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature AtomNet =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type of atom sets that can be efficiently matched and unified. *)
-(* ------------------------------------------------------------------------- *)
-
-type parameters = {fifo : bool}
-
-type 'a atomNet
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val new : parameters -> 'a atomNet
-
-val size : 'a atomNet -> int
-
-val insert : 'a atomNet -> Atom.atom * 'a -> 'a atomNet
-
-val fromList : parameters -> (Atom.atom * 'a) list -> 'a atomNet
-
-val filter : ('a -> bool) -> 'a atomNet -> 'a atomNet
-
-val toString : 'a atomNet -> string
-
-val pp : 'a Parser.pp -> 'a atomNet Parser.pp
-
-(* ------------------------------------------------------------------------- *)
-(* Matching and unification queries. *)
-(* *)
-(* These function return OVER-APPROXIMATIONS! *)
-(* Filter afterwards to get the precise set of satisfying values. *)
-(* ------------------------------------------------------------------------- *)
-
-val match : 'a atomNet -> Atom.atom -> 'a list
-
-val matched : 'a atomNet -> Atom.atom -> 'a list
-
-val unify : 'a atomNet -> Atom.atom -> 'a list
-
-end
--- a/src/Tools/Metis/src/AtomNet.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,59 +0,0 @@
-(* ========================================================================= *)
-(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure AtomNet :> AtomNet =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions. *)
-(* ------------------------------------------------------------------------- *)
-
-fun atomToTerm atom = Term.Fn atom;
-
-fun termToAtom (Term.Var _) = raise Bug "AtomNet.termToAtom"
- | termToAtom (Term.Fn atom) = atom;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of atom sets that can be efficiently matched and unified. *)
-(* ------------------------------------------------------------------------- *)
-
-type parameters = TermNet.parameters;
-
-type 'a atomNet = 'a TermNet.termNet;
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val new = TermNet.new;
-
-val size = TermNet.size;
-
-fun insert net (atm,a) = TermNet.insert net (atomToTerm atm, a);
-
-fun fromList parm l = foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
-
-val filter = TermNet.filter;
-
-fun toString net = "AtomNet[" ^ Int.toString (size net) ^ "]";
-
-val pp = TermNet.pp;
-
-(* ------------------------------------------------------------------------- *)
-(* Matching and unification queries. *)
-(* *)
-(* These function return OVER-APPROXIMATIONS! *)
-(* Filter afterwards to get the precise set of satisfying values. *)
-(* ------------------------------------------------------------------------- *)
-
-fun match net atm = TermNet.match net (atomToTerm atm);
-
-fun matched net atm = TermNet.matched net (atomToTerm atm);
-
-fun unify net atm = TermNet.unify net (atomToTerm atm);
-
-end
--- a/src/Tools/Metis/src/Clause.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,105 +0,0 @@
-(* ========================================================================= *)
-(* CLAUSE = ID + THEOREM *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Clause =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type of clause. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype literalOrder =
- NoLiteralOrder
- | UnsignedLiteralOrder
- | PositiveLiteralOrder
-
-type parameters =
- {ordering : KnuthBendixOrder.kbo,
- orderLiterals : literalOrder,
- orderTerms : bool}
-
-type clauseId = int
-
-type clauseInfo = {parameters : parameters, id : clauseId, thm : Thm.thm}
-
-type clause
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val default : parameters
-
-val newId : unit -> clauseId
-
-val mk : clauseInfo -> clause
-
-val dest : clause -> clauseInfo
-
-val id : clause -> clauseId
-
-val thm : clause -> Thm.thm
-
-val equalThms : clause -> clause -> bool
-
-val literals : clause -> Thm.clause
-
-val isTautology : clause -> bool
-
-val isContradiction : clause -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* The term ordering is used to cut down inferences. *)
-(* ------------------------------------------------------------------------- *)
-
-val largestLiterals : clause -> LiteralSet.set
-
-val largestEquations :
- clause -> (Literal.literal * Rewrite.orient * Term.term) list
-
-val largestSubterms :
- clause -> (Literal.literal * Term.path * Term.term) list
-
-val allSubterms : clause -> (Literal.literal * Term.path * Term.term) list
-
-(* ------------------------------------------------------------------------- *)
-(* Subsumption. *)
-(* ------------------------------------------------------------------------- *)
-
-val subsumes : clause Subsume.subsume -> clause -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* Simplifying rules: these preserve the clause id. *)
-(* ------------------------------------------------------------------------- *)
-
-val freshVars : clause -> clause
-
-val simplify : clause -> clause option
-
-val reduce : Units.units -> clause -> clause
-
-val rewrite : Rewrite.rewrite -> clause -> clause
-
-(* ------------------------------------------------------------------------- *)
-(* Inference rules: these generate new clause ids. *)
-(* ------------------------------------------------------------------------- *)
-
-val factor : clause -> clause list
-
-val resolve : clause * Literal.literal -> clause * Literal.literal -> clause
-
-val paramodulate :
- clause * Literal.literal * Rewrite.orient * Term.term ->
- clause * Literal.literal * Term.path * Term.term -> clause
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val showId : bool ref
-
-val pp : clause Parser.pp
-
-end
--- a/src/Tools/Metis/src/Clause.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,351 +0,0 @@
-(* ========================================================================= *)
-(* CLAUSE = ID + THEOREM *)
-(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Clause :> Clause =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions. *)
-(* ------------------------------------------------------------------------- *)
-
-val newId =
- let
- val r = ref 0
- in
- fn () => CRITICAL (fn () =>
- case r of ref n => let val () = r := n + 1 in n end)
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of clause. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype literalOrder =
- NoLiteralOrder
- | UnsignedLiteralOrder
- | PositiveLiteralOrder;
-
-type parameters =
- {ordering : KnuthBendixOrder.kbo,
- orderLiterals : literalOrder,
- orderTerms : bool};
-
-type clauseId = int;
-
-type clauseInfo = {parameters : parameters, id : clauseId, thm : Thm.thm};
-
-datatype clause = Clause of clauseInfo;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val showId = ref false;
-
-local
- val ppIdThm = Parser.ppPair Parser.ppInt Thm.pp;
-in
- fun pp p (Clause {id,thm,...}) =
- if !showId then ppIdThm p (id,thm) else Thm.pp p thm;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val default : parameters =
- {ordering = KnuthBendixOrder.default,
- orderLiterals = UnsignedLiteralOrder, (*LCP: changed from PositiveLiteralOrder*)
- orderTerms = true};
-
-fun mk info = Clause info
-
-fun dest (Clause info) = info;
-
-fun id (Clause {id = i, ...}) = i;
-
-fun thm (Clause {thm = th, ...}) = th;
-
-fun equalThms cl cl' = Thm.equal (thm cl) (thm cl');
-
-fun new parameters thm =
- Clause {parameters = parameters, id = newId (), thm = thm};
-
-fun literals cl = Thm.clause (thm cl);
-
-fun isTautology (Clause {thm,...}) = Thm.isTautology thm;
-
-fun isContradiction (Clause {thm,...}) = Thm.isContradiction thm;
-
-(* ------------------------------------------------------------------------- *)
-(* The term ordering is used to cut down inferences. *)
-(* ------------------------------------------------------------------------- *)
-
-fun strictlyLess ordering x_y =
- case KnuthBendixOrder.compare ordering x_y of
- SOME LESS => true
- | _ => false;
-
-fun isLargerTerm ({ordering,orderTerms,...} : parameters) l_r =
- not orderTerms orelse not (strictlyLess ordering l_r);
-
-local
- fun atomToTerms atm =
- case total Atom.destEq atm of
- NONE => [Term.Fn atm]
- | SOME (l,r) => [l,r];
-
- fun notStrictlyLess ordering (xs,ys) =
- let
- fun less x = List.exists (fn y => strictlyLess ordering (x,y)) ys
- in
- not (List.all less xs)
- end;
-in
- fun isLargerLiteral ({ordering,orderLiterals,...} : parameters) lits =
- case orderLiterals of
- NoLiteralOrder => K true
- | UnsignedLiteralOrder =>
- let
- fun addLit ((_,atm),acc) = atomToTerms atm @ acc
-
- val tms = LiteralSet.foldl addLit [] lits
- in
- fn (_,atm') => notStrictlyLess ordering (atomToTerms atm', tms)
- end
- | PositiveLiteralOrder =>
- case LiteralSet.findl (K true) lits of
- NONE => K true
- | SOME (pol,_) =>
- let
- fun addLit ((p,atm),acc) =
- if p = pol then atomToTerms atm @ acc else acc
-
- val tms = LiteralSet.foldl addLit [] lits
- in
- fn (pol',atm') =>
- if pol <> pol' then pol
- else notStrictlyLess ordering (atomToTerms atm', tms)
- end;
-end;
-
-fun largestLiterals (Clause {parameters,thm,...}) =
- let
- val litSet = Thm.clause thm
- val isLarger = isLargerLiteral parameters litSet
- fun addLit (lit,s) = if isLarger lit then LiteralSet.add s lit else s
- in
- LiteralSet.foldr addLit LiteralSet.empty litSet
- end;
-
-(*TRACE6
-val largestLiterals = fn cl =>
- let
- val ppResult = LiteralSet.pp
- val () = Parser.ppTrace pp "Clause.largestLiterals: cl" cl
- val result = largestLiterals cl
- val () = Parser.ppTrace ppResult "Clause.largestLiterals: result" result
- in
- result
- end;
-*)
-
-fun largestEquations (cl as Clause {parameters,...}) =
- let
- fun addEq lit ort (l_r as (l,_)) acc =
- if isLargerTerm parameters l_r then (lit,ort,l) :: acc else acc
-
- fun addLit (lit,acc) =
- case total Literal.destEq lit of
- NONE => acc
- | SOME (l,r) =>
- let
- val acc = addEq lit Rewrite.RightToLeft (r,l) acc
- val acc = addEq lit Rewrite.LeftToRight (l,r) acc
- in
- acc
- end
- in
- LiteralSet.foldr addLit [] (largestLiterals cl)
- end;
-
-local
- fun addLit (lit,acc) =
- let
- fun addTm ((path,tm),acc) = (lit,path,tm) :: acc
- in
- foldl addTm acc (Literal.nonVarTypedSubterms lit)
- end;
-in
- fun largestSubterms cl = LiteralSet.foldl addLit [] (largestLiterals cl);
-
- fun allSubterms cl = LiteralSet.foldl addLit [] (literals cl);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Subsumption. *)
-(* ------------------------------------------------------------------------- *)
-
-fun subsumes (subs : clause Subsume.subsume) cl =
- Subsume.isStrictlySubsumed subs (literals cl);
-
-(* ------------------------------------------------------------------------- *)
-(* Simplifying rules: these preserve the clause id. *)
-(* ------------------------------------------------------------------------- *)
-
-fun freshVars (Clause {parameters,id,thm}) =
- Clause {parameters = parameters, id = id, thm = Rule.freshVars thm};
-
-fun simplify (Clause {parameters,id,thm}) =
- case Rule.simplify thm of
- NONE => NONE
- | SOME thm => SOME (Clause {parameters = parameters, id = id, thm = thm});
-
-fun reduce units (Clause {parameters,id,thm}) =
- Clause {parameters = parameters, id = id, thm = Units.reduce units thm};
-
-fun rewrite rewr (cl as Clause {parameters,id,thm}) =
- let
- fun simp th =
- let
- val {ordering,...} = parameters
- val cmp = KnuthBendixOrder.compare ordering
- in
- Rewrite.rewriteIdRule rewr cmp id th
- end
-
-(*TRACE4
- val () = Parser.ppTrace Rewrite.pp "Clause.rewrite: rewr" rewr
- val () = Parser.ppTrace Parser.ppInt "Clause.rewrite: id" id
- val () = Parser.ppTrace pp "Clause.rewrite: cl" cl
-*)
-
- val thm =
- case Rewrite.peek rewr id of
- NONE => simp thm
- | SOME ((_,thm),_) => if Rewrite.isReduced rewr then thm else simp thm
-
- val result = Clause {parameters = parameters, id = id, thm = thm}
-
-(*TRACE4
- val () = Parser.ppTrace pp "Clause.rewrite: result" result
-*)
- in
- result
- end
-(*DEBUG
- handle Error err => raise Error ("Clause.rewrite:\n" ^ err);
-*)
-
-(* ------------------------------------------------------------------------- *)
-(* Inference rules: these generate new clause ids. *)
-(* ------------------------------------------------------------------------- *)
-
-fun factor (cl as Clause {parameters,thm,...}) =
- let
- val lits = largestLiterals cl
-
- fun apply sub = new parameters (Thm.subst sub thm)
- in
- map apply (Rule.factor' lits)
- end;
-
-(*TRACE5
-val factor = fn cl =>
- let
- val () = Parser.ppTrace pp "Clause.factor: cl" cl
- val result = factor cl
- val () = Parser.ppTrace (Parser.ppList pp) "Clause.factor: result" result
- in
- result
- end;
-*)
-
-fun resolve (cl1,lit1) (cl2,lit2) =
- let
-(*TRACE5
- val () = Parser.ppTrace pp "Clause.resolve: cl1" cl1
- val () = Parser.ppTrace Literal.pp "Clause.resolve: lit1" lit1
- val () = Parser.ppTrace pp "Clause.resolve: cl2" cl2
- val () = Parser.ppTrace Literal.pp "Clause.resolve: lit2" lit2
-*)
- val Clause {parameters, thm = th1, ...} = cl1
- and Clause {thm = th2, ...} = cl2
- val sub = Literal.unify Subst.empty lit1 (Literal.negate lit2)
-(*TRACE5
- val () = Parser.ppTrace Subst.pp "Clause.resolve: sub" sub
-*)
- val lit1 = Literal.subst sub lit1
- val lit2 = Literal.negate lit1
- val th1 = Thm.subst sub th1
- and th2 = Thm.subst sub th2
- val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse
-(*TRACE5
- (trace "Clause.resolve: th1 violates ordering\n"; false) orelse
-*)
- raise Error "resolve: clause1: ordering constraints"
- val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse
-(*TRACE5
- (trace "Clause.resolve: th2 violates ordering\n"; false) orelse
-*)
- raise Error "resolve: clause2: ordering constraints"
- val th = Thm.resolve lit1 th1 th2
-(*TRACE5
- val () = Parser.ppTrace Thm.pp "Clause.resolve: th" th
-*)
- val cl = Clause {parameters = parameters, id = newId (), thm = th}
-(*TRACE5
- val () = Parser.ppTrace pp "Clause.resolve: cl" cl
-*)
- in
- cl
- end;
-
-fun paramodulate (cl1,lit1,ort,tm1) (cl2,lit2,path,tm2) =
- let
-(*TRACE5
- val () = Parser.ppTrace pp "Clause.paramodulate: cl1" cl1
- val () = Parser.ppTrace Literal.pp "Clause.paramodulate: lit1" lit1
- val () = Parser.ppTrace pp "Clause.paramodulate: cl2" cl2
- val () = Parser.ppTrace Literal.pp "Clause.paramodulate: lit2" lit2
-*)
- val Clause {parameters, thm = th1, ...} = cl1
- and Clause {thm = th2, ...} = cl2
- val sub = Subst.unify Subst.empty tm1 tm2
- val lit1 = Literal.subst sub lit1
- and lit2 = Literal.subst sub lit2
- and th1 = Thm.subst sub th1
- and th2 = Thm.subst sub th2
- val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse
-(*TRACE5
- (trace "Clause.paramodulate: cl1 ordering\n"; false) orelse
-*)
- raise Error "paramodulate: with clause: ordering constraints"
- val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse
-(*TRACE5
- (trace "Clause.paramodulate: cl2 ordering\n"; false) orelse
-*)
- raise Error "paramodulate: into clause: ordering constraints"
- val eqn = (Literal.destEq lit1, th1)
- val eqn as (l_r,_) =
- case ort of
- Rewrite.LeftToRight => eqn
- | Rewrite.RightToLeft => Rule.symEqn eqn
- val _ = isLargerTerm parameters l_r orelse
-(*TRACE5
- (trace "Clause.paramodulate: eqn ordering\n"; false) orelse
-*)
- raise Error "paramodulate: equation: ordering constraints"
- val th = Rule.rewrRule eqn lit2 path th2
-(*TRACE5
- val () = Parser.ppTrace Thm.pp "Clause.paramodulate: th" th
-*)
- in
- Clause {parameters = parameters, id = newId (), thm = th}
- end;
-
-end
--- a/src/Tools/Metis/src/ElementSet.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,113 +0,0 @@
-(* ========================================================================= *)
-(* FINITE SETS WITH A FIXED ELEMENT TYPE *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature ElementSet =
-sig
-
-type element
-
-(* ------------------------------------------------------------------------- *)
-(* Finite sets *)
-(* ------------------------------------------------------------------------- *)
-
-type set
-
-val empty : set
-
-val singleton : element -> set
-
-val null : set -> bool
-
-val size : set -> int
-
-val member : element -> set -> bool
-
-val add : set -> element -> set
-
-val addList : set -> element list -> set
-
-val delete : set -> element -> set (* raises Error *)
-
-(* Union and intersect prefer elements in the second set *)
-
-val union : set -> set -> set
-
-val unionList : set list -> set
-
-val intersect : set -> set -> set
-
-val intersectList : set list -> set
-
-val difference : set -> set -> set
-
-val symmetricDifference : set -> set -> set
-
-val disjoint : set -> set -> bool
-
-val subset : set -> set -> bool
-
-val equal : set -> set -> bool
-
-val filter : (element -> bool) -> set -> set
-
-val partition : (element -> bool) -> set -> set * set
-
-val count : (element -> bool) -> set -> int
-
-val foldl : (element * 's -> 's) -> 's -> set -> 's
-
-val foldr : (element * 's -> 's) -> 's -> set -> 's
-
-val findl : (element -> bool) -> set -> element option
-
-val findr : (element -> bool) -> set -> element option
-
-val firstl : (element -> 'a option) -> set -> 'a option
-
-val firstr : (element -> 'a option) -> set -> 'a option
-
-val exists : (element -> bool) -> set -> bool
-
-val all : (element -> bool) -> set -> bool
-
-val map : (element -> 'a) -> set -> (element * 'a) list
-
-val transform : (element -> 'a) -> set -> 'a list
-
-val app : (element -> unit) -> set -> unit
-
-val toList : set -> element list
-
-val fromList : element list -> set
-
-val pick : set -> element (* raises Empty *)
-
-val random : set -> element (* raises Empty *)
-
-val deletePick : set -> element * set (* raises Empty *)
-
-val deleteRandom : set -> element * set (* raises Empty *)
-
-val compare : set * set -> order
-
-val close : (set -> set) -> set -> set
-
-val toString : set -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators over sets *)
-(* ------------------------------------------------------------------------- *)
-
-type iterator
-
-val mkIterator : set -> iterator option
-
-val mkRevIterator : set -> iterator option
-
-val readIterator : iterator -> element
-
-val advanceIterator : iterator -> iterator option
-
-end
--- a/src/Tools/Metis/src/ElementSet.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,123 +0,0 @@
-(* ========================================================================= *)
-(* FINITE SETS WITH A FIXED ELEMENT TYPE *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-functor ElementSet (Key : Ordered) :> ElementSet where type element = Key.t =
-struct
-
-(*isabelle open Metis;*)
-
-type element = Key.t;
-
-(* ------------------------------------------------------------------------- *)
-(* Finite sets *)
-(* ------------------------------------------------------------------------- *)
-
-type set = Key.t Set.set;
-
-val empty = Set.empty Key.compare;
-
-fun singleton key = Set.singleton Key.compare key;
-
-val null = Set.null;
-
-val size = Set.size;
-
-val member = Set.member;
-
-val add = Set.add;
-
-val addList = Set.addList;
-
-val delete = Set.delete;
-
-val union = Set.union;
-
-val unionList = Set.unionList;
-
-val intersect = Set.intersect;
-
-val intersectList = Set.intersectList;
-
-val difference = Set.difference;
-
-val symmetricDifference = Set.symmetricDifference;
-
-val disjoint = Set.disjoint;
-
-val subset = Set.subset;
-
-val equal = Set.equal;
-
-val filter = Set.filter;
-
-val partition = Set.partition;
-
-val count = Set.count;
-
-val foldl = Set.foldl;
-
-val foldr = Set.foldr;
-
-val findl = Set.findl;
-
-val findr = Set.findr;
-
-val firstl = Set.firstl;
-
-val firstr = Set.firstr;
-
-val exists = Set.exists;
-
-val all = Set.all;
-
-val map = Set.map;
-
-val transform = Set.transform;
-
-val app = Set.app;
-
-val toList = Set.toList;
-
-fun fromList l = Set.fromList Key.compare l;
-
-val pick = Set.pick;
-
-val random = Set.random;
-
-val deletePick = Set.deletePick;
-
-val deleteRandom = Set.deleteRandom;
-
-val compare = Set.compare;
-
-val close = Set.close;
-
-val toString = Set.toString;
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators over sets *)
-(* ------------------------------------------------------------------------- *)
-
-type iterator = Key.t Set.iterator;
-
-val mkIterator = Set.mkIterator;
-
-val mkRevIterator = Set.mkRevIterator;
-
-val readIterator = Set.readIterator;
-
-val advanceIterator = Set.advanceIterator;
-
-end
-
-(*isabelle structure Metis = struct open Metis;*)
-
-structure IntSet =
-ElementSet (IntOrdered);
-
-structure StringSet =
-ElementSet (StringOrdered);
-
-(*isabelle end;*)
--- a/src/Tools/Metis/src/FILES Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-Portable.sig PortableIsabelle.sml
-PP.sig PP.sml
-Useful.sig Useful.sml
-Lazy.sig Lazy.sml
-Ordered.sig Ordered.sml
-Set.sig RandomSet.sml Set.sml
-ElementSet.sig ElementSet.sml
-Map.sig RandomMap.sml Map.sml
-KeyMap.sig KeyMap.sml
-Sharing.sig Sharing.sml
-Stream.sig Stream.sml
-Heap.sig Heap.sml
-Parser.sig Parser.sml
-Options.sig Options.sml
-Name.sig Name.sml
-Term.sig Term.sml
-Subst.sig Subst.sml
-Atom.sig Atom.sml
-Formula.sig Formula.sml
-Literal.sig Literal.sml
-Thm.sig Thm.sml
-Proof.sig Proof.sml
-Rule.sig Rule.sml
-Normalize.sig Normalize.sml
-Model.sig Model.sml
-Problem.sig Problem.sml
-TermNet.sig TermNet.sml
-AtomNet.sig AtomNet.sml
-LiteralNet.sig LiteralNet.sml
-Subsume.sig Subsume.sml
-KnuthBendixOrder.sig KnuthBendixOrder.sml
-Rewrite.sig Rewrite.sml
-Units.sig Units.sml
-Clause.sig Clause.sml
-Active.sig Active.sml
-Waiting.sig Waiting.sml
-Resolution.sig Resolution.sml
-Tptp.sig Tptp.sml
--- a/src/Tools/Metis/src/Formula.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,179 +0,0 @@
-(* ========================================================================= *)
-(* FIRST ORDER LOGIC FORMULAS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Formula =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type of first order logic formulas. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formula =
- True
- | False
- | Atom of Atom.atom
- | Not of formula
- | And of formula * formula
- | Or of formula * formula
- | Imp of formula * formula
- | Iff of formula * formula
- | Forall of Term.var * formula
- | Exists of Term.var * formula
-
-(* ------------------------------------------------------------------------- *)
-(* Constructors and destructors. *)
-(* ------------------------------------------------------------------------- *)
-
-(* Booleans *)
-
-val mkBoolean : bool -> formula
-
-val destBoolean : formula -> bool
-
-val isBoolean : formula -> bool
-
-(* Functions *)
-
-val functions : formula -> NameAritySet.set
-
-val functionNames : formula -> NameSet.set
-
-(* Relations *)
-
-val relations : formula -> NameAritySet.set
-
-val relationNames : formula -> NameSet.set
-
-(* Atoms *)
-
-val destAtom : formula -> Atom.atom
-
-val isAtom : formula -> bool
-
-(* Negations *)
-
-val destNeg : formula -> formula
-
-val isNeg : formula -> bool
-
-val stripNeg : formula -> int * formula
-
-(* Conjunctions *)
-
-val listMkConj : formula list -> formula
-
-val stripConj : formula -> formula list
-
-val flattenConj : formula -> formula list
-
-(* Disjunctions *)
-
-val listMkDisj : formula list -> formula
-
-val stripDisj : formula -> formula list
-
-val flattenDisj : formula -> formula list
-
-(* Equivalences *)
-
-val listMkEquiv : formula list -> formula
-
-val stripEquiv : formula -> formula list
-
-val flattenEquiv : formula -> formula list
-
-(* Universal quantification *)
-
-val destForall : formula -> Term.var * formula
-
-val isForall : formula -> bool
-
-val listMkForall : Term.var list * formula -> formula
-
-val stripForall : formula -> Term.var list * formula
-
-(* Existential quantification *)
-
-val destExists : formula -> Term.var * formula
-
-val isExists : formula -> bool
-
-val listMkExists : Term.var list * formula -> formula
-
-val stripExists : formula -> Term.var list * formula
-
-(* ------------------------------------------------------------------------- *)
-(* The size of a formula in symbols. *)
-(* ------------------------------------------------------------------------- *)
-
-val symbols : formula -> int
-
-(* ------------------------------------------------------------------------- *)
-(* A total comparison function for formulas. *)
-(* ------------------------------------------------------------------------- *)
-
-val compare : formula * formula -> order
-
-(* ------------------------------------------------------------------------- *)
-(* Free variables. *)
-(* ------------------------------------------------------------------------- *)
-
-val freeIn : Term.var -> formula -> bool
-
-val freeVars : formula -> NameSet.set
-
-val specialize : formula -> formula
-
-val generalize : formula -> formula
-
-(* ------------------------------------------------------------------------- *)
-(* Substitutions. *)
-(* ------------------------------------------------------------------------- *)
-
-val subst : Subst.subst -> formula -> formula
-
-(* ------------------------------------------------------------------------- *)
-(* The equality relation. *)
-(* ------------------------------------------------------------------------- *)
-
-val mkEq : Term.term * Term.term -> formula
-
-val destEq : formula -> Term.term * Term.term
-
-val isEq : formula -> bool
-
-val mkNeq : Term.term * Term.term -> formula
-
-val destNeq : formula -> Term.term * Term.term
-
-val isNeq : formula -> bool
-
-val mkRefl : Term.term -> formula
-
-val destRefl : formula -> Term.term
-
-val isRefl : formula -> bool
-
-val sym : formula -> formula (* raises Error if given a refl *)
-
-val lhs : formula -> Term.term
-
-val rhs : formula -> Term.term
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty-printing. *)
-(* ------------------------------------------------------------------------- *)
-
-type quotation = formula Parser.quotation
-
-val pp : formula Parser.pp
-
-val toString : formula -> string
-
-val fromString : string -> formula
-
-val parse : quotation -> formula
-
-end
--- a/src/Tools/Metis/src/Formula.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,515 +0,0 @@
-(* ========================================================================= *)
-(* FIRST ORDER LOGIC FORMULAS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Formula :> Formula =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of first order logic formulas. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formula =
- True
- | False
- | Atom of Atom.atom
- | Not of formula
- | And of formula * formula
- | Or of formula * formula
- | Imp of formula * formula
- | Iff of formula * formula
- | Forall of Term.var * formula
- | Exists of Term.var * formula;
-
-(* ------------------------------------------------------------------------- *)
-(* Constructors and destructors. *)
-(* ------------------------------------------------------------------------- *)
-
-(* Booleans *)
-
-fun mkBoolean true = True
- | mkBoolean false = False;
-
-fun destBoolean True = true
- | destBoolean False = false
- | destBoolean _ = raise Error "destBoolean";
-
-val isBoolean = can destBoolean;
-
-(* Functions *)
-
-local
- fun funcs fs [] = fs
- | funcs fs (True :: fms) = funcs fs fms
- | funcs fs (False :: fms) = funcs fs fms
- | funcs fs (Atom atm :: fms) =
- funcs (NameAritySet.union (Atom.functions atm) fs) fms
- | funcs fs (Not p :: fms) = funcs fs (p :: fms)
- | funcs fs (And (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Or (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Imp (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Iff (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Forall (_,p) :: fms) = funcs fs (p :: fms)
- | funcs fs (Exists (_,p) :: fms) = funcs fs (p :: fms);
-in
- fun functions fm = funcs NameAritySet.empty [fm];
-end;
-
-local
- fun funcs fs [] = fs
- | funcs fs (True :: fms) = funcs fs fms
- | funcs fs (False :: fms) = funcs fs fms
- | funcs fs (Atom atm :: fms) =
- funcs (NameSet.union (Atom.functionNames atm) fs) fms
- | funcs fs (Not p :: fms) = funcs fs (p :: fms)
- | funcs fs (And (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Or (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Imp (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Iff (p,q) :: fms) = funcs fs (p :: q :: fms)
- | funcs fs (Forall (_,p) :: fms) = funcs fs (p :: fms)
- | funcs fs (Exists (_,p) :: fms) = funcs fs (p :: fms);
-in
- fun functionNames fm = funcs NameSet.empty [fm];
-end;
-
-(* Relations *)
-
-local
- fun rels fs [] = fs
- | rels fs (True :: fms) = rels fs fms
- | rels fs (False :: fms) = rels fs fms
- | rels fs (Atom atm :: fms) =
- rels (NameAritySet.add fs (Atom.relation atm)) fms
- | rels fs (Not p :: fms) = rels fs (p :: fms)
- | rels fs (And (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Or (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Imp (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Iff (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Forall (_,p) :: fms) = rels fs (p :: fms)
- | rels fs (Exists (_,p) :: fms) = rels fs (p :: fms);
-in
- fun relations fm = rels NameAritySet.empty [fm];
-end;
-
-local
- fun rels fs [] = fs
- | rels fs (True :: fms) = rels fs fms
- | rels fs (False :: fms) = rels fs fms
- | rels fs (Atom atm :: fms) = rels (NameSet.add fs (Atom.name atm)) fms
- | rels fs (Not p :: fms) = rels fs (p :: fms)
- | rels fs (And (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Or (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Imp (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Iff (p,q) :: fms) = rels fs (p :: q :: fms)
- | rels fs (Forall (_,p) :: fms) = rels fs (p :: fms)
- | rels fs (Exists (_,p) :: fms) = rels fs (p :: fms);
-in
- fun relationNames fm = rels NameSet.empty [fm];
-end;
-
-(* Atoms *)
-
-fun destAtom (Atom atm) = atm
- | destAtom _ = raise Error "Formula.destAtom";
-
-val isAtom = can destAtom;
-
-(* Negations *)
-
-fun destNeg (Not p) = p
- | destNeg _ = raise Error "Formula.destNeg";
-
-val isNeg = can destNeg;
-
-val stripNeg =
- let
- fun strip n (Not fm) = strip (n + 1) fm
- | strip n fm = (n,fm)
- in
- strip 0
- end;
-
-(* Conjunctions *)
-
-fun listMkConj fms =
- case rev fms of [] => True | fm :: fms => foldl And fm fms;
-
-local
- fun strip cs (And (p,q)) = strip (p :: cs) q
- | strip cs fm = rev (fm :: cs);
-in
- fun stripConj True = []
- | stripConj fm = strip [] fm;
-end;
-
-val flattenConj =
- let
- fun flat acc [] = acc
- | flat acc (And (p,q) :: fms) = flat acc (q :: p :: fms)
- | flat acc (True :: fms) = flat acc fms
- | flat acc (fm :: fms) = flat (fm :: acc) fms
- in
- fn fm => flat [] [fm]
- end;
-
-(* Disjunctions *)
-
-fun listMkDisj fms =
- case rev fms of [] => False | fm :: fms => foldl Or fm fms;
-
-local
- fun strip cs (Or (p,q)) = strip (p :: cs) q
- | strip cs fm = rev (fm :: cs);
-in
- fun stripDisj False = []
- | stripDisj fm = strip [] fm;
-end;
-
-val flattenDisj =
- let
- fun flat acc [] = acc
- | flat acc (Or (p,q) :: fms) = flat acc (q :: p :: fms)
- | flat acc (False :: fms) = flat acc fms
- | flat acc (fm :: fms) = flat (fm :: acc) fms
- in
- fn fm => flat [] [fm]
- end;
-
-(* Equivalences *)
-
-fun listMkEquiv fms =
- case rev fms of [] => True | fm :: fms => foldl Iff fm fms;
-
-local
- fun strip cs (Iff (p,q)) = strip (p :: cs) q
- | strip cs fm = rev (fm :: cs);
-in
- fun stripEquiv True = []
- | stripEquiv fm = strip [] fm;
-end;
-
-val flattenEquiv =
- let
- fun flat acc [] = acc
- | flat acc (Iff (p,q) :: fms) = flat acc (q :: p :: fms)
- | flat acc (True :: fms) = flat acc fms
- | flat acc (fm :: fms) = flat (fm :: acc) fms
- in
- fn fm => flat [] [fm]
- end;
-
-(* Universal quantifiers *)
-
-fun destForall (Forall v_f) = v_f
- | destForall _ = raise Error "destForall";
-
-val isForall = can destForall;
-
-fun listMkForall ([],body) = body
- | listMkForall (v :: vs, body) = Forall (v, listMkForall (vs,body));
-
-local
- fun strip vs (Forall (v,b)) = strip (v :: vs) b
- | strip vs tm = (rev vs, tm);
-in
- val stripForall = strip [];
-end;
-
-(* Existential quantifiers *)
-
-fun destExists (Exists v_f) = v_f
- | destExists _ = raise Error "destExists";
-
-val isExists = can destExists;
-
-fun listMkExists ([],body) = body
- | listMkExists (v :: vs, body) = Exists (v, listMkExists (vs,body));
-
-local
- fun strip vs (Exists (v,b)) = strip (v :: vs) b
- | strip vs tm = (rev vs, tm);
-in
- val stripExists = strip [];
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* The size of a formula in symbols. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun sz n [] = n
- | sz n (True :: fms) = sz (n + 1) fms
- | sz n (False :: fms) = sz (n + 1) fms
- | sz n (Atom atm :: fms) = sz (n + Atom.symbols atm) fms
- | sz n (Not p :: fms) = sz (n + 1) (p :: fms)
- | sz n (And (p,q) :: fms) = sz (n + 1) (p :: q :: fms)
- | sz n (Or (p,q) :: fms) = sz (n + 1) (p :: q :: fms)
- | sz n (Imp (p,q) :: fms) = sz (n + 1) (p :: q :: fms)
- | sz n (Iff (p,q) :: fms) = sz (n + 1) (p :: q :: fms)
- | sz n (Forall (_,p) :: fms) = sz (n + 1) (p :: fms)
- | sz n (Exists (_,p) :: fms) = sz (n + 1) (p :: fms);
-in
- fun symbols fm = sz 0 [fm];
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* A total comparison function for formulas. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun cmp [] = EQUAL
- | cmp ((True,True) :: l) = cmp l
- | cmp ((True,_) :: _) = LESS
- | cmp ((_,True) :: _) = GREATER
- | cmp ((False,False) :: l) = cmp l
- | cmp ((False,_) :: _) = LESS
- | cmp ((_,False) :: _) = GREATER
- | cmp ((Atom atm1, Atom atm2) :: l) =
- (case Atom.compare (atm1,atm2) of
- LESS => LESS
- | EQUAL => cmp l
- | GREATER => GREATER)
- | cmp ((Atom _, _) :: _) = LESS
- | cmp ((_, Atom _) :: _) = GREATER
- | cmp ((Not p1, Not p2) :: l) = cmp ((p1,p2) :: l)
- | cmp ((Not _, _) :: _) = LESS
- | cmp ((_, Not _) :: _) = GREATER
- | cmp ((And (p1,q1), And (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l)
- | cmp ((And _, _) :: _) = LESS
- | cmp ((_, And _) :: _) = GREATER
- | cmp ((Or (p1,q1), Or (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l)
- | cmp ((Or _, _) :: _) = LESS
- | cmp ((_, Or _) :: _) = GREATER
- | cmp ((Imp (p1,q1), Imp (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l)
- | cmp ((Imp _, _) :: _) = LESS
- | cmp ((_, Imp _) :: _) = GREATER
- | cmp ((Iff (p1,q1), Iff (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l)
- | cmp ((Iff _, _) :: _) = LESS
- | cmp ((_, Iff _) :: _) = GREATER
- | cmp ((Forall (v1,p1), Forall (v2,p2)) :: l) =
- (case Name.compare (v1,v2) of
- LESS => LESS
- | EQUAL => cmp ((p1,p2) :: l)
- | GREATER => GREATER)
- | cmp ((Forall _, Exists _) :: _) = LESS
- | cmp ((Exists _, Forall _) :: _) = GREATER
- | cmp ((Exists (v1,p1), Exists (v2,p2)) :: l) =
- (case Name.compare (v1,v2) of
- LESS => LESS
- | EQUAL => cmp ((p1,p2) :: l)
- | GREATER => GREATER);
-in
- fun compare fm1_fm2 = cmp [fm1_fm2];
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Free variables. *)
-(* ------------------------------------------------------------------------- *)
-
-fun freeIn v =
- let
- fun f [] = false
- | f (True :: fms) = f fms
- | f (False :: fms) = f fms
- | f (Atom atm :: fms) = Atom.freeIn v atm orelse f fms
- | f (Not p :: fms) = f (p :: fms)
- | f (And (p,q) :: fms) = f (p :: q :: fms)
- | f (Or (p,q) :: fms) = f (p :: q :: fms)
- | f (Imp (p,q) :: fms) = f (p :: q :: fms)
- | f (Iff (p,q) :: fms) = f (p :: q :: fms)
- | f (Forall (w,p) :: fms) = if v = w then f fms else f (p :: fms)
- | f (Exists (w,p) :: fms) = if v = w then f fms else f (p :: fms)
- in
- fn fm => f [fm]
- end;
-
-local
- fun fv vs [] = vs
- | fv vs ((_,True) :: fms) = fv vs fms
- | fv vs ((_,False) :: fms) = fv vs fms
- | fv vs ((bv, Atom atm) :: fms) =
- fv (NameSet.union vs (NameSet.difference (Atom.freeVars atm) bv)) fms
- | fv vs ((bv, Not p) :: fms) = fv vs ((bv,p) :: fms)
- | fv vs ((bv, And (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
- | fv vs ((bv, Or (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
- | fv vs ((bv, Imp (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
- | fv vs ((bv, Iff (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
- | fv vs ((bv, Forall (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms)
- | fv vs ((bv, Exists (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms);
-in
- fun freeVars fm = fv NameSet.empty [(NameSet.empty,fm)];
-end;
-
-fun specialize fm = snd (stripForall fm);
-
-fun generalize fm = listMkForall (NameSet.toList (freeVars fm), fm);
-
-(* ------------------------------------------------------------------------- *)
-(* Substitutions. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun substCheck sub fm = if Subst.null sub then fm else substFm sub fm
-
- and substFm sub fm =
- case fm of
- True => fm
- | False => fm
- | Atom (p,tms) =>
- let
- val tms' = Sharing.map (Subst.subst sub) tms
- in
- if Sharing.pointerEqual (tms,tms') then fm else Atom (p,tms')
- end
- | Not p =>
- let
- val p' = substFm sub p
- in
- if Sharing.pointerEqual (p,p') then fm else Not p'
- end
- | And (p,q) => substConn sub fm And p q
- | Or (p,q) => substConn sub fm Or p q
- | Imp (p,q) => substConn sub fm Imp p q
- | Iff (p,q) => substConn sub fm Iff p q
- | Forall (v,p) => substQuant sub fm Forall v p
- | Exists (v,p) => substQuant sub fm Exists v p
-
- and substConn sub fm conn p q =
- let
- val p' = substFm sub p
- and q' = substFm sub q
- in
- if Sharing.pointerEqual (p,p') andalso
- Sharing.pointerEqual (q,q')
- then fm
- else conn (p',q')
- end
-
- and substQuant sub fm quant v p =
- let
- val v' =
- let
- fun f (w,s) =
- if w = v then s
- else
- case Subst.peek sub w of
- NONE => NameSet.add s w
- | SOME tm => NameSet.union s (Term.freeVars tm)
-
- val vars = freeVars p
- val vars = NameSet.foldl f NameSet.empty vars
- in
- Term.variantPrime vars v
- end
-
- val sub =
- if v = v' then Subst.remove sub (NameSet.singleton v)
- else Subst.insert sub (v, Term.Var v')
-
- val p' = substCheck sub p
- in
- if v = v' andalso Sharing.pointerEqual (p,p') then fm
- else quant (v',p')
- end;
-in
- val subst = substCheck;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* The equality relation. *)
-(* ------------------------------------------------------------------------- *)
-
-fun mkEq a_b = Atom (Atom.mkEq a_b);
-
-fun destEq fm = Atom.destEq (destAtom fm);
-
-val isEq = can destEq;
-
-fun mkNeq a_b = Not (mkEq a_b);
-
-fun destNeq (Not fm) = destEq fm
- | destNeq _ = raise Error "Formula.destNeq";
-
-val isNeq = can destNeq;
-
-fun mkRefl tm = Atom (Atom.mkRefl tm);
-
-fun destRefl fm = Atom.destRefl (destAtom fm);
-
-val isRefl = can destRefl;
-
-fun sym fm = Atom (Atom.sym (destAtom fm));
-
-fun lhs fm = fst (destEq fm);
-
-fun rhs fm = snd (destEq fm);
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty-printing. *)
-(* ------------------------------------------------------------------------- *)
-
-type quotation = formula Parser.quotation
-
-val truthSymbol = "T"
-and falsitySymbol = "F"
-and conjunctionSymbol = "/\\"
-and disjunctionSymbol = "\\/"
-and implicationSymbol = "==>"
-and equivalenceSymbol = "<=>"
-and universalSymbol = "!"
-and existentialSymbol = "?";
-
-local
- fun demote True = Term.Fn (truthSymbol,[])
- | demote False = Term.Fn (falsitySymbol,[])
- | demote (Atom (p,tms)) = Term.Fn (p,tms)
- | demote (Not p) = Term.Fn (!Term.negation, [demote p])
- | demote (And (p,q)) = Term.Fn (conjunctionSymbol, [demote p, demote q])
- | demote (Or (p,q)) = Term.Fn (disjunctionSymbol, [demote p, demote q])
- | demote (Imp (p,q)) = Term.Fn (implicationSymbol, [demote p, demote q])
- | demote (Iff (p,q)) = Term.Fn (equivalenceSymbol, [demote p, demote q])
- | demote (Forall (v,b)) = Term.Fn (universalSymbol, [Term.Var v, demote b])
- | demote (Exists (v,b)) =
- Term.Fn (existentialSymbol, [Term.Var v, demote b]);
-in
- fun pp ppstrm fm = Term.pp ppstrm (demote fm);
-end;
-
-val toString = Parser.toString pp;
-
-local
- fun isQuant [Term.Var _, _] = true
- | isQuant _ = false;
-
- fun promote (Term.Var v) = Atom (v,[])
- | promote (Term.Fn (f,tms)) =
- if f = truthSymbol andalso null tms then
- True
- else if f = falsitySymbol andalso null tms then
- False
- else if f = !Term.negation andalso length tms = 1 then
- Not (promote (hd tms))
- else if f = conjunctionSymbol andalso length tms = 2 then
- And (promote (hd tms), promote (List.nth (tms,1)))
- else if f = disjunctionSymbol andalso length tms = 2 then
- Or (promote (hd tms), promote (List.nth (tms,1)))
- else if f = implicationSymbol andalso length tms = 2 then
- Imp (promote (hd tms), promote (List.nth (tms,1)))
- else if f = equivalenceSymbol andalso length tms = 2 then
- Iff (promote (hd tms), promote (List.nth (tms,1)))
- else if f = universalSymbol andalso isQuant tms then
- Forall (Term.destVar (hd tms), promote (List.nth (tms,1)))
- else if f = existentialSymbol andalso isQuant tms then
- Exists (Term.destVar (hd tms), promote (List.nth (tms,1)))
- else
- Atom (f,tms);
-in
- fun fromString s = promote (Term.fromString s);
-end;
-
-val parse = Parser.parseQuotation toString fromString;
-
-end
--- a/src/Tools/Metis/src/Heap.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,31 +0,0 @@
-(* ========================================================================= *)
-(* A HEAP DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Heap =
-sig
-
-type 'a heap
-
-val new : ('a * 'a -> order) -> 'a heap
-
-val add : 'a heap -> 'a -> 'a heap
-
-val null : 'a heap -> bool
-
-val top : 'a heap -> 'a (* raises Empty *)
-
-val remove : 'a heap -> 'a * 'a heap (* raises Empty *)
-
-val size : 'a heap -> int
-
-val app : ('a -> unit) -> 'a heap -> unit
-
-val toList : 'a heap -> 'a list
-
-val toStream : 'a heap -> 'a Stream.stream
-
-val toString : 'a heap -> string
-
-end
--- a/src/Tools/Metis/src/Heap.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,77 +0,0 @@
-(* ========================================================================= *)
-(* A HEAP DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Heap :> Heap =
-struct
-
-(* Leftist heaps as in Purely Functional Data Structures, by Chris Okasaki *)
-
-datatype 'a node = E | T of int * 'a * 'a node * 'a node;
-
-datatype 'a heap = Heap of ('a * 'a -> order) * int * 'a node;
-
-fun rank E = 0
- | rank (T (r,_,_,_)) = r;
-
-fun makeT (x,a,b) =
- if rank a >= rank b then T (rank b + 1, x, a, b) else T (rank a + 1, x, b, a);
-
-fun merge cmp =
- let
- fun mrg (h,E) = h
- | mrg (E,h) = h
- | mrg (h1 as T (_,x,a1,b1), h2 as T (_,y,a2,b2)) =
- case cmp (x,y) of
- GREATER => makeT (y, a2, mrg (h1,b2))
- | _ => makeT (x, a1, mrg (b1,h2))
- in
- mrg
- end;
-
-fun new cmp = Heap (cmp,0,E);
-
-fun add (Heap (f,n,a)) x = Heap (f, n + 1, merge f (T (1,x,E,E), a));
-
-fun size (Heap (_, n, _)) = n;
-
-fun null h = size h = 0;
-
-fun top (Heap (_,_,E)) = raise Empty
- | top (Heap (_, _, T (_,x,_,_))) = x;
-
-fun remove (Heap (_,_,E)) = raise Empty
- | remove (Heap (f, n, T (_,x,a,b))) = (x, Heap (f, n - 1, merge f (a,b)));
-
-fun app f =
- let
- fun ap [] = ()
- | ap (E :: rest) = ap rest
- | ap (T (_,d,a,b) :: rest) = (f d; ap (a :: b :: rest))
- in
- fn Heap (_,_,a) => ap [a]
- end;
-
-fun toList h =
- if null h then []
- else
- let
- val (x,h) = remove h
- in
- x :: toList h
- end;
-
-fun toStream h =
- if null h then Stream.NIL
- else
- let
- val (x,h) = remove h
- in
- Stream.CONS (x, fn () => toStream h)
- end;
-
-fun toString h =
- "Heap[" ^ (if null h then "" else Int.toString (size h)) ^ "]";
-
-end
--- a/src/Tools/Metis/src/KeyMap.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,103 +0,0 @@
-(* ========================================================================= *)
-(* FINITE MAPS WITH A FIXED KEY TYPE *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature KeyMap =
-sig
-
-type key
-
-(* ------------------------------------------------------------------------- *)
-(* Finite maps *)
-(* ------------------------------------------------------------------------- *)
-
-type 'a map
-
-val new : unit -> 'a map
-
-val null : 'a map -> bool
-
-val size : 'a map -> int
-
-val singleton : key * 'a -> 'a map
-
-val inDomain : key -> 'a map -> bool
-
-val peek : 'a map -> key -> 'a option
-
-val insert : 'a map -> key * 'a -> 'a map
-
-val insertList : 'a map -> (key * 'a) list -> 'a map
-
-val get : 'a map -> key -> 'a (* raises Error *)
-
-(* Union and intersect prefer keys in the second map *)
-
-val union : ('a * 'a -> 'a option) -> 'a map -> 'a map -> 'a map
-
-val intersect : ('a * 'a -> 'a option) -> 'a map -> 'a map -> 'a map
-
-val delete : 'a map -> key -> 'a map (* raises Error *)
-
-val difference : 'a map -> 'a map -> 'a map
-
-val subsetDomain : 'a map -> 'a map -> bool
-
-val equalDomain : 'a map -> 'a map -> bool
-
-val mapPartial : (key * 'a -> 'b option) -> 'a map -> 'b map
-
-val filter : (key * 'a -> bool) -> 'a map -> 'a map
-
-val map : (key * 'a -> 'b) -> 'a map -> 'b map
-
-val app : (key * 'a -> unit) -> 'a map -> unit
-
-val transform : ('a -> 'b) -> 'a map -> 'b map
-
-val foldl : (key * 'a * 's -> 's) -> 's -> 'a map -> 's
-
-val foldr : (key * 'a * 's -> 's) -> 's -> 'a map -> 's
-
-val findl : (key * 'a -> bool) -> 'a map -> (key * 'a) option
-
-val findr : (key * 'a -> bool) -> 'a map -> (key * 'a) option
-
-val firstl : (key * 'a -> 'b option) -> 'a map -> 'b option
-
-val firstr : (key * 'a -> 'b option) -> 'a map -> 'b option
-
-val exists : (key * 'a -> bool) -> 'a map -> bool
-
-val all : (key * 'a -> bool) -> 'a map -> bool
-
-val domain : 'a map -> key list
-
-val toList : 'a map -> (key * 'a) list
-
-val fromList : (key * 'a) list -> 'a map
-
-val compare : ('a * 'a -> order) -> 'a map * 'a map -> order
-
-val equal : ('a -> 'a -> bool) -> 'a map -> 'a map -> bool
-
-val random : 'a map -> key * 'a (* raises Empty *)
-
-val toString : 'a map -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators over maps *)
-(* ------------------------------------------------------------------------- *)
-
-type 'a iterator
-
-val mkIterator : 'a map -> 'a iterator option
-
-val mkRevIterator : 'a map -> 'a iterator option
-
-val readIterator : 'a iterator -> key * 'a
-
-val advanceIterator : 'a iterator -> 'a iterator option
-
-end
--- a/src/Tools/Metis/src/KeyMap.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-(* ========================================================================= *)
-(* FINITE MAPS WITH A FIXED KEY TYPE *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-functor KeyMap (Key : Ordered) :> KeyMap where type key = Key.t =
-struct
-
-(*isabelle open Metis;*)
-
-type key = Key.t;
-
-(* ------------------------------------------------------------------------- *)
-(* Finite maps *)
-(* ------------------------------------------------------------------------- *)
-
-type 'a map = (Key.t,'a) Map.map;
-
-fun new () = Map.new Key.compare;
-
-val null = Map.null;
-
-val size = Map.size;
-
-fun singleton key_value = Map.singleton Key.compare key_value;
-
-val inDomain = Map.inDomain;
-
-val peek = Map.peek;
-
-val insert = Map.insert;
-
-val insertList = Map.insertList;
-
-val get = Map.get;
-
-(* Both union and intersect prefer keys in the second map *)
-
-val union = Map.union;
-
-val intersect = Map.intersect;
-
-val delete = Map.delete;
-
-val difference = Map.difference;
-
-val subsetDomain = Map.subsetDomain;
-
-val equalDomain = Map.equalDomain;
-
-val mapPartial = Map.mapPartial;
-
-val filter = Map.filter;
-
-val map = Map.map;
-
-val app = Map.app;
-
-val transform = Map.transform;
-
-val foldl = Map.foldl;
-
-val foldr = Map.foldr;
-
-val findl = Map.findl;
-
-val findr = Map.findr;
-
-val firstl = Map.firstl;
-
-val firstr = Map.firstr;
-
-val exists = Map.exists;
-
-val all = Map.all;
-
-val domain = Map.domain;
-
-val toList = Map.toList;
-
-fun fromList l = Map.fromList Key.compare l;
-
-val compare = Map.compare;
-
-val equal = Map.equal;
-
-val random = Map.random;
-
-val toString = Map.toString;
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators over maps *)
-(* ------------------------------------------------------------------------- *)
-
-type 'a iterator = (Key.t,'a) Map.iterator;
-
-val mkIterator = Map.mkIterator;
-
-val mkRevIterator = Map.mkRevIterator;
-
-val readIterator = Map.readIterator;
-
-val advanceIterator = Map.advanceIterator;
-
-end
-
-(*isabelle structure Metis = struct open Metis*)
-
-structure IntMap =
-KeyMap (IntOrdered);
-
-structure StringMap =
-KeyMap (StringOrdered);
-
-(*isabelle end;*)
--- a/src/Tools/Metis/src/KnuthBendixOrder.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,22 +0,0 @@
-(* ========================================================================= *)
-(* THE KNUTH-BENDIX TERM ORDERING *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature KnuthBendixOrder =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* The weight of all constants must be at least 1, and there must be at most *)
-(* one unary function with weight 0. *)
-(* ------------------------------------------------------------------------- *)
-
-type kbo =
- {weight : Term.function -> int,
- precedence : Term.function * Term.function -> order}
-
-val default : kbo
-
-val compare : kbo -> Term.term * Term.term -> order option
-
-end
--- a/src/Tools/Metis/src/KnuthBendixOrder.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,251 +0,0 @@
-(* ========================================================================= *)
-(* KNUTH-BENDIX TERM ORDERING CONSTRAINTS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure KnuthBendixOrder :> KnuthBendixOrder =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions. *)
-(* ------------------------------------------------------------------------- *)
-
-fun firstNotEqual f l =
- case List.find op<> l of
- SOME (x,y) => f x y
- | NONE => raise Bug "firstNotEqual";
-
-(* ------------------------------------------------------------------------- *)
-(* The weight of all constants must be at least 1, and there must be at most *)
-(* one unary function with weight 0. *)
-(* ------------------------------------------------------------------------- *)
-
-type kbo =
- {weight : Term.function -> int,
- precedence : Term.function * Term.function -> order};
-
-(* Default weight = uniform *)
-
-val uniformWeight : Term.function -> int = K 1;
-
-(* Default precedence = by arity *)
-
-val arityPrecedence : Term.function * Term.function -> order =
- fn ((f1,n1),(f2,n2)) =>
- case Int.compare (n1,n2) of
- LESS => LESS
- | EQUAL => String.compare (f1,f2)
- | GREATER => GREATER;
-
-(* The default order *)
-
-val default = {weight = uniformWeight, precedence = arityPrecedence};
-
-(* ------------------------------------------------------------------------- *)
-(* Term weight-1 represented as a linear function of the weight-1 of the *)
-(* variables in the term (plus a constant). *)
-(* *)
-(* Note that the conditions on weight functions ensure that all weights are *)
-(* at least 1, so all weight-1s are at least 0. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype weight = Weight of int NameMap.map * int;
-
-val weightEmpty : int NameMap.map = NameMap.new ();
-
-val weightZero = Weight (weightEmpty,0);
-
-fun weightIsZero (Weight (m,c)) = c = 0 andalso NameMap.null m;
-
-fun weightNeg (Weight (m,c)) = Weight (NameMap.transform ~ m, ~c);
-
-local
- fun add (n1,n2) =
- let
- val n = n1 + n2
- in
- if n = 0 then NONE else SOME n
- end;
-in
- fun weightAdd (Weight (m1,c1)) (Weight (m2,c2)) =
- Weight (NameMap.union add m1 m2, c1 + c2);
-end;
-
-fun weightSubtract w1 w2 = weightAdd w1 (weightNeg w2);
-
-fun weightMult 0 _ = weightZero
- | weightMult 1 w = w
- | weightMult k (Weight (m,c)) =
- let
- fun mult n = k * n
- in
- Weight (NameMap.transform mult m, k * c)
- end;
-
-fun weightTerm weight =
- let
- fun wt m c [] = Weight (m,c)
- | wt m c (Term.Var v :: tms) =
- let
- val n = Option.getOpt (NameMap.peek m v, 0)
- in
- wt (NameMap.insert m (v, n + 1)) (c + 1) tms
- end
- | wt m c (Term.Fn (f,a) :: tms) =
- wt m (c + weight (f, length a)) (a @ tms)
- in
- fn tm => wt weightEmpty ~1 [tm]
- end;
-
-fun weightIsVar v (Weight (m,c)) =
- c = 0 andalso NameMap.size m = 1 andalso NameMap.peek m v = SOME 1;
-
-fun weightConst (Weight (_,c)) = c;
-
-fun weightVars (Weight (m,_)) =
- NameMap.foldl (fn (v,_,s) => NameSet.add s v) NameSet.empty m;
-
-val weightsVars =
- List.foldl (fn (w,s) => NameSet.union (weightVars w) s) NameSet.empty;
-
-fun weightVarList w = NameSet.toList (weightVars w);
-
-fun weightNumVars (Weight (m,_)) = NameMap.size m;
-
-fun weightNumVarsWithPositiveCoeff (Weight (m,_)) =
- NameMap.foldl (fn (_,n,z) => if n > 0 then z + 1 else z) 0 m;
-
-fun weightCoeff var (Weight (m,_)) = Option.getOpt (NameMap.peek m var, 0);
-
-fun weightCoeffs varList w = map (fn var => weightCoeff var w) varList;
-
-fun weightCoeffSum (Weight (m,_)) = NameMap.foldl (fn (_,n,z) => n + z) 0 m;
-
-fun weightLowerBound (w as Weight (m,c)) =
- if NameMap.exists (fn (_,n) => n < 0) m then NONE else SOME c;
-
-fun weightNoLowerBound w = not (Option.isSome (weightLowerBound w));
-
-fun weightAlwaysPositive w =
- case weightLowerBound w of NONE => false | SOME n => n > 0;
-
-fun weightUpperBound (w as Weight (m,c)) =
- if NameMap.exists (fn (_,n) => n > 0) m then NONE else SOME c;
-
-fun weightNoUpperBound w = not (Option.isSome (weightUpperBound w));
-
-fun weightAlwaysNegative w =
- case weightUpperBound w of NONE => false | SOME n => n < 0;
-
-fun weightGcd (w as Weight (m,c)) =
- NameMap.foldl (fn (_,i,k) => gcd i k) (Int.abs c) m;
-
-fun ppWeightList pp =
- let
- fun coeffToString n =
- if n < 0 then "~" ^ coeffToString (~n)
- else if n = 1 then ""
- else Int.toString n
-
- fun pp_tm pp ("",n) = Parser.ppInt pp n
- | pp_tm pp (v,n) = Parser.ppString pp (coeffToString n ^ v)
- in
- fn [] => Parser.ppInt pp 0
- | tms => Parser.ppSequence " +" pp_tm pp tms
- end;
-
-fun ppWeight pp (Weight (m,c)) =
- let
- val l = NameMap.toList m
- val l = if c = 0 then l else l @ [("",c)]
- in
- ppWeightList pp l
- end;
-
-val weightToString = Parser.toString ppWeight;
-
-(* ------------------------------------------------------------------------- *)
-(* The Knuth-Bendix term order. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype kboOrder = Less | Equal | Greater | SignOf of weight;
-
-fun kboOrder {weight,precedence} =
- let
- fun weightDifference tm1 tm2 =
- let
- val w1 = weightTerm weight tm1
- and w2 = weightTerm weight tm2
- in
- weightSubtract w2 w1
- end
-
- fun weightLess tm1 tm2 =
- let
- val w = weightDifference tm1 tm2
- in
- if weightIsZero w then precedenceLess tm1 tm2
- else weightDiffLess w tm1 tm2
- end
-
- and weightDiffLess w tm1 tm2 =
- case weightLowerBound w of
- NONE => false
- | SOME 0 => precedenceLess tm1 tm2
- | SOME n => n > 0
-
- and precedenceLess (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) =
- (case precedence ((f1, length a1), (f2, length a2)) of
- LESS => true
- | EQUAL => firstNotEqual weightLess (zip a1 a2)
- | GREATER => false)
- | precedenceLess _ _ = false
-
- fun weightDiffGreater w tm1 tm2 = weightDiffLess (weightNeg w) tm2 tm1
-
- fun weightCmp tm1 tm2 =
- let
- val w = weightDifference tm1 tm2
- in
- if weightIsZero w then precedenceCmp tm1 tm2
- else if weightDiffLess w tm1 tm2 then Less
- else if weightDiffGreater w tm1 tm2 then Greater
- else SignOf w
- end
-
- and precedenceCmp (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) =
- (case precedence ((f1, length a1), (f2, length a2)) of
- LESS => Less
- | EQUAL => firstNotEqual weightCmp (zip a1 a2)
- | GREATER => Greater)
- | precedenceCmp _ _ = raise Bug "kboOrder.precendenceCmp"
- in
- fn (tm1,tm2) => if tm1 = tm2 then Equal else weightCmp tm1 tm2
- end;
-
-fun compare kbo tm1_tm2 =
- case kboOrder kbo tm1_tm2 of
- Less => SOME LESS
- | Equal => SOME EQUAL
- | Greater => SOME GREATER
- | SignOf _ => NONE;
-
-(*TRACE7
-val compare = fn kbo => fn (tm1,tm2) =>
- let
- val () = Parser.ppTrace Term.pp "KnuthBendixOrder.compare: tm1" tm1
- val () = Parser.ppTrace Term.pp "KnuthBendixOrder.compare: tm2" tm2
- val result = compare kbo (tm1,tm2)
- val () =
- case result of
- NONE => trace "KnuthBendixOrder.compare: result = Incomparable\n"
- | SOME x =>
- Parser.ppTrace Parser.ppOrder "KnuthBendixOrder.compare: result" x
- in
- result
- end;
-*)
-
-end
--- a/src/Tools/Metis/src/Lazy.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-(* ========================================================================= *)
-(* SUPPORT FOR LAZY EVALUATION *)
-(* Copyright (c) 2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Lazy =
-sig
-
-type 'a lazy
-
-val delay : (unit -> 'a) -> 'a lazy
-
-val force : 'a lazy -> 'a
-
-val memoize : (unit -> 'a) -> unit -> 'a
-
-end
--- a/src/Tools/Metis/src/Lazy.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-(* ========================================================================= *)
-(* SUPPORT FOR LAZY EVALUATION *)
-(* Copyright (c) 2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Lazy :> Lazy =
-struct
-
-datatype 'a thunk =
- Value of 'a
- | Thunk of unit -> 'a;
-
-datatype 'a lazy = Lazy of 'a thunk ref;
-
-fun delay f = Lazy (ref (Thunk f));
-
-fun force (Lazy (ref (Value v))) = v
- | force (Lazy (s as ref (Thunk f))) =
- let
- val v = f ()
- val () = s := Value v
- in
- v
- end;
-
-fun memoize f =
- let
- val t = delay f
- in
- fn () => force t
- end;
-
-end
--- a/src/Tools/Metis/src/Literal.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,163 +0,0 @@
-(* ========================================================================= *)
-(* FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Literal =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type for storing first order logic literals. *)
-(* ------------------------------------------------------------------------- *)
-
-type polarity = bool
-
-type literal = polarity * Atom.atom
-
-(* ------------------------------------------------------------------------- *)
-(* Constructors and destructors. *)
-(* ------------------------------------------------------------------------- *)
-
-val polarity : literal -> polarity
-
-val atom : literal -> Atom.atom
-
-val name : literal -> Atom.relationName
-
-val arguments : literal -> Term.term list
-
-val arity : literal -> int
-
-val positive : literal -> bool
-
-val negative : literal -> bool
-
-val negate : literal -> literal
-
-val relation : literal -> Atom.relation
-
-val functions : literal -> NameAritySet.set
-
-val functionNames : literal -> NameSet.set
-
-(* Binary relations *)
-
-val mkBinop : Atom.relationName -> polarity * Term.term * Term.term -> literal
-
-val destBinop : Atom.relationName -> literal -> polarity * Term.term * Term.term
-
-val isBinop : Atom.relationName -> literal -> bool
-
-(* Formulas *)
-
-val toFormula : literal -> Formula.formula
-
-val fromFormula : Formula.formula -> literal
-
-(* ------------------------------------------------------------------------- *)
-(* The size of a literal in symbols. *)
-(* ------------------------------------------------------------------------- *)
-
-val symbols : literal -> int
-
-(* ------------------------------------------------------------------------- *)
-(* A total comparison function for literals. *)
-(* ------------------------------------------------------------------------- *)
-
-val compare : literal * literal -> order (* negative < positive *)
-
-(* ------------------------------------------------------------------------- *)
-(* Subterms. *)
-(* ------------------------------------------------------------------------- *)
-
-val subterm : literal -> Term.path -> Term.term
-
-val subterms : literal -> (Term.path * Term.term) list
-
-val replace : literal -> Term.path * Term.term -> literal
-
-(* ------------------------------------------------------------------------- *)
-(* Free variables. *)
-(* ------------------------------------------------------------------------- *)
-
-val freeIn : Term.var -> literal -> bool
-
-val freeVars : literal -> NameSet.set
-
-(* ------------------------------------------------------------------------- *)
-(* Substitutions. *)
-(* ------------------------------------------------------------------------- *)
-
-val subst : Subst.subst -> literal -> literal
-
-(* ------------------------------------------------------------------------- *)
-(* Matching. *)
-(* ------------------------------------------------------------------------- *)
-
-val match : (* raises Error *)
- Subst.subst -> literal -> literal -> Subst.subst
-
-(* ------------------------------------------------------------------------- *)
-(* Unification. *)
-(* ------------------------------------------------------------------------- *)
-
-val unify : (* raises Error *)
- Subst.subst -> literal -> literal -> Subst.subst
-
-(* ------------------------------------------------------------------------- *)
-(* The equality relation. *)
-(* ------------------------------------------------------------------------- *)
-
-val mkEq : Term.term * Term.term -> literal
-
-val destEq : literal -> Term.term * Term.term
-
-val isEq : literal -> bool
-
-val mkNeq : Term.term * Term.term -> literal
-
-val destNeq : literal -> Term.term * Term.term
-
-val isNeq : literal -> bool
-
-val mkRefl : Term.term -> literal
-
-val destRefl : literal -> Term.term
-
-val isRefl : literal -> bool
-
-val mkIrrefl : Term.term -> literal
-
-val destIrrefl : literal -> Term.term
-
-val isIrrefl : literal -> bool
-
-(* The following work with both equalities and disequalities *)
-
-val sym : literal -> literal (* raises Error if given a refl or irrefl *)
-
-val lhs : literal -> Term.term
-
-val rhs : literal -> Term.term
-
-(* ------------------------------------------------------------------------- *)
-(* Special support for terms with type annotations. *)
-(* ------------------------------------------------------------------------- *)
-
-val typedSymbols : literal -> int
-
-val nonVarTypedSubterms : literal -> (Term.path * Term.term) list
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty-printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val pp : literal Parser.pp
-
-val toString : literal -> string
-
-val fromString : string -> literal
-
-val parse : Term.term Parser.quotation -> literal
-
-end
--- a/src/Tools/Metis/src/Literal.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,273 +0,0 @@
-(* ========================================================================= *)
-(* FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Literal :> Literal =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* A type for storing first order logic literals. *)
-(* ------------------------------------------------------------------------- *)
-
-type polarity = bool;
-
-type literal = polarity * Atom.atom;
-
-(* ------------------------------------------------------------------------- *)
-(* Constructors and destructors. *)
-(* ------------------------------------------------------------------------- *)
-
-fun polarity ((pol,_) : literal) = pol;
-
-fun atom ((_,atm) : literal) = atm;
-
-fun name lit = Atom.name (atom lit);
-
-fun arguments lit = Atom.arguments (atom lit);
-
-fun arity lit = Atom.arity (atom lit);
-
-fun positive lit = polarity lit;
-
-fun negative lit = not (polarity lit);
-
-fun negate (pol,atm) : literal = (not pol, atm)
-
-fun relation lit = Atom.relation (atom lit);
-
-fun functions lit = Atom.functions (atom lit);
-
-fun functionNames lit = Atom.functionNames (atom lit);
-
-(* Binary relations *)
-
-fun mkBinop rel (pol,a,b) : literal = (pol, Atom.mkBinop rel (a,b));
-
-fun destBinop rel ((pol,atm) : literal) =
- case Atom.destBinop rel atm of (a,b) => (pol,a,b);
-
-fun isBinop rel = can (destBinop rel);
-
-(* Formulas *)
-
-fun toFormula (true,atm) = Formula.Atom atm
- | toFormula (false,atm) = Formula.Not (Formula.Atom atm);
-
-fun fromFormula (Formula.Atom atm) = (true,atm)
- | fromFormula (Formula.Not (Formula.Atom atm)) = (false,atm)
- | fromFormula _ = raise Error "Literal.fromFormula";
-
-(* ------------------------------------------------------------------------- *)
-(* The size of a literal in symbols. *)
-(* ------------------------------------------------------------------------- *)
-
-fun symbols ((_,atm) : literal) = Atom.symbols atm;
-
-(* ------------------------------------------------------------------------- *)
-(* A total comparison function for literals. *)
-(* ------------------------------------------------------------------------- *)
-
-fun compare ((pol1,atm1),(pol2,atm2)) =
- case boolCompare (pol1,pol2) of
- LESS => GREATER
- | EQUAL => Atom.compare (atm1,atm2)
- | GREATER => LESS;
-
-(* ------------------------------------------------------------------------- *)
-(* Subterms. *)
-(* ------------------------------------------------------------------------- *)
-
-fun subterm lit path = Atom.subterm (atom lit) path;
-
-fun subterms lit = Atom.subterms (atom lit);
-
-fun replace (lit as (pol,atm)) path_tm =
- let
- val atm' = Atom.replace atm path_tm
- in
- if Sharing.pointerEqual (atm,atm') then lit else (pol,atm')
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Free variables. *)
-(* ------------------------------------------------------------------------- *)
-
-fun freeIn v lit = Atom.freeIn v (atom lit);
-
-fun freeVars lit = Atom.freeVars (atom lit);
-
-(* ------------------------------------------------------------------------- *)
-(* Substitutions. *)
-(* ------------------------------------------------------------------------- *)
-
-fun subst sub (lit as (pol,atm)) : literal =
- let
- val atm' = Atom.subst sub atm
- in
- if Sharing.pointerEqual (atm',atm) then lit else (pol,atm')
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Matching. *)
-(* ------------------------------------------------------------------------- *)
-
-fun match sub ((pol1,atm1) : literal) (pol2,atm2) =
- let
- val _ = pol1 = pol2 orelse raise Error "Literal.match"
- in
- Atom.match sub atm1 atm2
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Unification. *)
-(* ------------------------------------------------------------------------- *)
-
-fun unify sub ((pol1,atm1) : literal) (pol2,atm2) =
- let
- val _ = pol1 = pol2 orelse raise Error "Literal.unify"
- in
- Atom.unify sub atm1 atm2
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* The equality relation. *)
-(* ------------------------------------------------------------------------- *)
-
-fun mkEq l_r : literal = (true, Atom.mkEq l_r);
-
-fun destEq ((true,atm) : literal) = Atom.destEq atm
- | destEq (false,_) = raise Error "Literal.destEq";
-
-val isEq = can destEq;
-
-fun mkNeq l_r : literal = (false, Atom.mkEq l_r);
-
-fun destNeq ((false,atm) : literal) = Atom.destEq atm
- | destNeq (true,_) = raise Error "Literal.destNeq";
-
-val isNeq = can destNeq;
-
-fun mkRefl tm = (true, Atom.mkRefl tm);
-
-fun destRefl (true,atm) = Atom.destRefl atm
- | destRefl (false,_) = raise Error "Literal.destRefl";
-
-val isRefl = can destRefl;
-
-fun mkIrrefl tm = (false, Atom.mkRefl tm);
-
-fun destIrrefl (true,_) = raise Error "Literal.destIrrefl"
- | destIrrefl (false,atm) = Atom.destRefl atm;
-
-val isIrrefl = can destIrrefl;
-
-fun sym (pol,atm) : literal = (pol, Atom.sym atm);
-
-fun lhs ((_,atm) : literal) = Atom.lhs atm;
-
-fun rhs ((_,atm) : literal) = Atom.rhs atm;
-
-(* ------------------------------------------------------------------------- *)
-(* Special support for terms with type annotations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun typedSymbols ((_,atm) : literal) = Atom.typedSymbols atm;
-
-fun nonVarTypedSubterms ((_,atm) : literal) = Atom.nonVarTypedSubterms atm;
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty-printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val pp = Parser.ppMap toFormula Formula.pp;
-
-val toString = Parser.toString pp;
-
-fun fromString s = fromFormula (Formula.fromString s);
-
-val parse = Parser.parseQuotation Term.toString fromString;
-
-end
-
-structure LiteralOrdered =
-struct type t = Literal.literal val compare = Literal.compare end
-
-structure LiteralSet =
-struct
-
- local
- structure S = ElementSet (LiteralOrdered);
- in
- open S;
- end;
-
- fun negateMember lit set = member (Literal.negate lit) set;
-
- val negate =
- let
- fun f (lit,set) = add set (Literal.negate lit)
- in
- foldl f empty
- end;
-
- val relations =
- let
- fun f (lit,set) = NameAritySet.add set (Literal.relation lit)
- in
- foldl f NameAritySet.empty
- end;
-
- val functions =
- let
- fun f (lit,set) = NameAritySet.union set (Literal.functions lit)
- in
- foldl f NameAritySet.empty
- end;
-
- val freeVars =
- let
- fun f (lit,set) = NameSet.union set (Literal.freeVars lit)
- in
- foldl f NameSet.empty
- end;
-
- val symbols =
- let
- fun f (lit,z) = Literal.symbols lit + z
- in
- foldl f 0
- end;
-
- val typedSymbols =
- let
- fun f (lit,z) = Literal.typedSymbols lit + z
- in
- foldl f 0
- end;
-
- fun subst sub lits =
- let
- fun substLit (lit,(eq,lits')) =
- let
- val lit' = Literal.subst sub lit
- val eq = eq andalso Sharing.pointerEqual (lit,lit')
- in
- (eq, add lits' lit')
- end
-
- val (eq,lits') = foldl substLit (true,empty) lits
- in
- if eq then lits else lits'
- end;
-
- val pp =
- Parser.ppMap
- toList
- (Parser.ppBracket "{" "}" (Parser.ppSequence "," Literal.pp));
-
-end
-
-structure LiteralMap = KeyMap (LiteralOrdered);
--- a/src/Tools/Metis/src/LiteralNet.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-(* ========================================================================= *)
-(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature LiteralNet =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type of literal sets that can be efficiently matched and unified. *)
-(* ------------------------------------------------------------------------- *)
-
-type parameters = {fifo : bool}
-
-type 'a literalNet
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-val new : parameters -> 'a literalNet
-
-val size : 'a literalNet -> int
-
-val profile : 'a literalNet -> {positive : int, negative : int}
-
-val insert : 'a literalNet -> Literal.literal * 'a -> 'a literalNet
-
-val fromList : parameters -> (Literal.literal * 'a) list -> 'a literalNet
-
-val filter : ('a -> bool) -> 'a literalNet -> 'a literalNet
-
-val toString : 'a literalNet -> string
-
-val pp : 'a Parser.pp -> 'a literalNet Parser.pp
-
-(* ------------------------------------------------------------------------- *)
-(* Matching and unification queries. *)
-(* *)
-(* These function return OVER-APPROXIMATIONS! *)
-(* Filter afterwards to get the precise set of satisfying values. *)
-(* ------------------------------------------------------------------------- *)
-
-val match : 'a literalNet -> Literal.literal -> 'a list
-
-val matched : 'a literalNet -> Literal.literal -> 'a list
-
-val unify : 'a literalNet -> Literal.literal -> 'a list
-
-end
--- a/src/Tools/Metis/src/LiteralNet.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,74 +0,0 @@
-(* ========================================================================= *)
-(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure LiteralNet :> LiteralNet =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of literal sets that can be efficiently matched and unified. *)
-(* ------------------------------------------------------------------------- *)
-
-type parameters = AtomNet.parameters;
-
-type 'a literalNet =
- {positive : 'a AtomNet.atomNet,
- negative : 'a AtomNet.atomNet};
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun new parm = {positive = AtomNet.new parm, negative = AtomNet.new parm};
-
-local
- fun pos ({positive,...} : 'a literalNet) = AtomNet.size positive;
-
- fun neg ({negative,...} : 'a literalNet) = AtomNet.size negative;
-in
- fun size net = pos net + neg net;
-
- fun profile net = {positive = pos net, negative = neg net};
-end;
-
-fun insert {positive,negative} ((true,atm),a) =
- {positive = AtomNet.insert positive (atm,a), negative = negative}
- | insert {positive,negative} ((false,atm),a) =
- {positive = positive, negative = AtomNet.insert negative (atm,a)};
-
-fun fromList parm l = foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
-
-fun filter pred {positive,negative} =
- {positive = AtomNet.filter pred positive,
- negative = AtomNet.filter pred negative};
-
-fun toString net = "LiteralNet[" ^ Int.toString (size net) ^ "]";
-
-fun pp ppA =
- Parser.ppMap
- (fn {positive,negative} => (positive,negative))
- (Parser.ppBinop " + NEGATIVE" (AtomNet.pp ppA) (AtomNet.pp ppA));
-
-(* ------------------------------------------------------------------------- *)
-(* Matching and unification queries. *)
-(* *)
-(* These function return OVER-APPROXIMATIONS! *)
-(* Filter afterwards to get the precise set of satisfying values. *)
-(* ------------------------------------------------------------------------- *)
-
-fun match ({positive,...} : 'a literalNet) (true,atm) =
- AtomNet.match positive atm
- | match {negative,...} (false,atm) = AtomNet.match negative atm;
-
-fun matched ({positive,...} : 'a literalNet) (true,atm) =
- AtomNet.matched positive atm
- | matched {negative,...} (false,atm) = AtomNet.matched negative atm;
-
-fun unify ({positive,...} : 'a literalNet) (true,atm) =
- AtomNet.unify positive atm
- | unify {negative,...} (false,atm) = AtomNet.unify negative atm;
-
-end
--- a/src/Tools/Metis/src/Map.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,103 +0,0 @@
-(* ========================================================================= *)
-(* FINITE MAPS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Map =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Finite maps *)
-(* ------------------------------------------------------------------------- *)
-
-type ('key,'a) map
-
-val new : ('key * 'key -> order) -> ('key,'a) map
-
-val null : ('key,'a) map -> bool
-
-val size : ('key,'a) map -> int
-
-val singleton : ('key * 'key -> order) -> 'key * 'a -> ('key,'a) map
-
-val inDomain : 'key -> ('key,'a) map -> bool
-
-val peek : ('key,'a) map -> 'key -> 'a option
-
-val insert : ('key,'a) map -> 'key * 'a -> ('key,'a) map
-
-val insertList : ('key,'a) map -> ('key * 'a) list -> ('key,'a) map
-
-val get : ('key,'a) map -> 'key -> 'a (* raises Error *)
-
-(* Union and intersect prefer keys in the second map *)
-
-val union :
- ('a * 'a -> 'a option) -> ('key,'a) map -> ('key,'a) map -> ('key,'a) map
-
-val intersect :
- ('a * 'a -> 'a option) -> ('key,'a) map -> ('key,'a) map -> ('key,'a) map
-
-val delete : ('key,'a) map -> 'key -> ('key,'a) map (* raises Error *)
-
-val difference : ('key,'a) map -> ('key,'b) map -> ('key,'a) map
-
-val subsetDomain : ('key,'a) map -> ('key,'a) map -> bool
-
-val equalDomain : ('key,'a) map -> ('key,'a) map -> bool
-
-val mapPartial : ('key * 'a -> 'b option) -> ('key,'a) map -> ('key,'b) map
-
-val filter : ('key * 'a -> bool) -> ('key,'a) map -> ('key,'a) map
-
-val map : ('key * 'a -> 'b) -> ('key,'a) map -> ('key,'b) map
-
-val app : ('key * 'a -> unit) -> ('key,'a) map -> unit
-
-val transform : ('a -> 'b) -> ('key,'a) map -> ('key,'b) map
-
-val foldl : ('key * 'a * 's -> 's) -> 's -> ('key,'a) map -> 's
-
-val foldr : ('key * 'a * 's -> 's) -> 's -> ('key,'a) map -> 's
-
-val findl : ('key * 'a -> bool) -> ('key,'a) map -> ('key * 'a) option
-
-val findr : ('key * 'a -> bool) -> ('key,'a) map -> ('key * 'a) option
-
-val firstl : ('key * 'a -> 'b option) -> ('key,'a) map -> 'b option
-
-val firstr : ('key * 'a -> 'b option) -> ('key,'a) map -> 'b option
-
-val exists : ('key * 'a -> bool) -> ('key,'a) map -> bool
-
-val all : ('key * 'a -> bool) -> ('key,'a) map -> bool
-
-val domain : ('key,'a) map -> 'key list
-
-val toList : ('key,'a) map -> ('key * 'a) list
-
-val fromList : ('key * 'key -> order) -> ('key * 'a) list -> ('key,'a) map
-
-val random : ('key,'a) map -> 'key * 'a (* raises Empty *)
-
-val compare : ('a * 'a -> order) -> ('key,'a) map * ('key,'a) map -> order
-
-val equal : ('a -> 'a -> bool) -> ('key,'a) map -> ('key,'a) map -> bool
-
-val toString : ('key,'a) map -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators over maps *)
-(* ------------------------------------------------------------------------- *)
-
-type ('key,'a) iterator
-
-val mkIterator : ('key,'a) map -> ('key,'a) iterator option
-
-val mkRevIterator : ('key,'a) map -> ('key,'a) iterator option
-
-val readIterator : ('key,'a) iterator -> 'key * 'a
-
-val advanceIterator : ('key,'a) iterator -> ('key,'a) iterator option
-
-end
--- a/src/Tools/Metis/src/Map.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-(* ========================================================================= *)
-(* FINITE MAPS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Map = RandomMap;
--- a/src/Tools/Metis/src/Model.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,94 +0,0 @@
-(* ========================================================================= *)
-(* RANDOM FINITE MODELS *)
-(* Copyright (c) 2003-2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Model =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* The parts of the model that are fixed. *)
-(* Note: a model of size N has integer elements 0...N-1. *)
-(* ------------------------------------------------------------------------- *)
-
-type fixed =
- {size : int} ->
- {functions : (Term.functionName * int list) -> int option,
- relations : (Atom.relationName * int list) -> bool option}
-
-val fixedMerge : fixed -> fixed -> fixed (* Prefers the second fixed *)
-
-val fixedMergeList : fixed list -> fixed
-
-val fixedPure : fixed (* : = *)
-
-val fixedBasic : fixed (* id fst snd #1 #2 #3 <> *)
-
-val fixedModulo : fixed (* <numerals> suc pre ~ + - * exp div mod *)
- (* is_0 divides even odd *)
-
-val fixedOverflowNum : fixed (* <numerals> suc pre + - * exp div mod *)
- (* is_0 <= < >= > divides even odd *)
-
-val fixedOverflowInt : fixed (* <numerals> suc pre + - * exp div mod *)
- (* is_0 <= < >= > divides even odd *)
-
-val fixedSet : fixed (* empty univ union intersect compl card in subset *)
-
-(* ------------------------------------------------------------------------- *)
-(* A type of random finite models. *)
-(* ------------------------------------------------------------------------- *)
-
-type parameters = {size : int, fixed : fixed}
-
-type model
-
-val new : parameters -> model
-
-val size : model -> int
-
-(* ------------------------------------------------------------------------- *)
-(* Valuations. *)
-(* ------------------------------------------------------------------------- *)
-
-type valuation = int NameMap.map
-
-val valuationEmpty : valuation
-
-val valuationRandom : {size : int} -> NameSet.set -> valuation
-
-val valuationFold :
- {size : int} -> NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a
-
-(* ------------------------------------------------------------------------- *)
-(* Interpreting terms and formulas in the model. *)
-(* ------------------------------------------------------------------------- *)
-
-val interpretTerm : model -> valuation -> Term.term -> int
-
-val interpretAtom : model -> valuation -> Atom.atom -> bool
-
-val interpretFormula : model -> valuation -> Formula.formula -> bool
-
-val interpretLiteral : model -> valuation -> Literal.literal -> bool
-
-val interpretClause : model -> valuation -> Thm.clause -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* Check whether random groundings of a formula are true in the model. *)
-(* Note: if it's cheaper, a systematic check will be performed instead. *)
-(* ------------------------------------------------------------------------- *)
-
-val checkAtom :
- {maxChecks : int} -> model -> Atom.atom -> {T : int, F : int}
-
-val checkFormula :
- {maxChecks : int} -> model -> Formula.formula -> {T : int, F : int}
-
-val checkLiteral :
- {maxChecks : int} -> model -> Literal.literal -> {T : int, F : int}
-
-val checkClause :
- {maxChecks : int} -> model -> Thm.clause -> {T : int, F : int}
-
-end
--- a/src/Tools/Metis/src/Model.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,591 +0,0 @@
-(* ========================================================================= *)
-(* RANDOM FINITE MODELS *)
-(* Copyright (c) 2003-2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Model :> Model =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions. *)
-(* ------------------------------------------------------------------------- *)
-
-fun intExp x y = exp op* x y 1;
-
-fun natFromString "" = NONE
- | natFromString "0" = SOME 0
- | natFromString s =
- case charToInt (String.sub (s,0)) of
- NONE => NONE
- | SOME 0 => NONE
- | SOME d =>
- let
- fun parse 0 _ acc = SOME acc
- | parse n i acc =
- case charToInt (String.sub (s,i)) of
- NONE => NONE
- | SOME d => parse (n - 1) (i + 1) (10 * acc + d)
- in
- parse (size s - 1) 1 d
- end;
-
-fun projection (_,[]) = NONE
- | projection ("#1", x :: _) = SOME x
- | projection ("#2", _ :: x :: _) = SOME x
- | projection ("#3", _ :: _ :: x :: _) = SOME x
- | projection (func,args) =
- let
- val f = size func
- and n = length args
-
- val p =
- if f < 2 orelse n <= 3 orelse String.sub (func,0) <> #"#" then NONE
- else if f = 2 then
- (case charToInt (String.sub (func,1)) of
- NONE => NONE
- | p as SOME d => if d <= 3 then NONE else p)
- else if (n < intExp 10 (f - 2) handle Overflow => true) then NONE
- else
- (natFromString (String.extract (func,1,NONE))
- handle Overflow => NONE)
- in
- case p of
- NONE => NONE
- | SOME k => if k > n then NONE else SOME (List.nth (args, k - 1))
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* The parts of the model that are fixed. *)
-(* Note: a model of size N has integer elements 0...N-1. *)
-(* ------------------------------------------------------------------------- *)
-
-type fixedModel =
- {functions : (Term.functionName * int list) -> int option,
- relations : (Atom.relationName * int list) -> bool option};
-
-type fixed = {size : int} -> fixedModel
-
-fun fixedMerge fixed1 fixed2 parm =
- let
- val {functions = f1, relations = r1} = fixed1 parm
- and {functions = f2, relations = r2} = fixed2 parm
-
- fun functions x = case f2 x of NONE => f1 x | s => s
-
- fun relations x = case r2 x of NONE => r1 x | s => s
- in
- {functions = functions, relations = relations}
- end;
-
-fun fixedMergeList [] = raise Bug "fixedMergeList: empty"
- | fixedMergeList (f :: l) = foldl (uncurry fixedMerge) f l;
-
-fun fixedPure {size = _} =
- let
- fun functions (":",[x,_]) = SOME x
- | functions _ = NONE
-
- fun relations (rel,[x,y]) =
- if (rel,2) = Atom.eqRelation then SOME (x = y) else NONE
- | relations _ = NONE
- in
- {functions = functions, relations = relations}
- end;
-
-fun fixedBasic {size = _} =
- let
- fun functions ("id",[x]) = SOME x
- | functions ("fst",[x,_]) = SOME x
- | functions ("snd",[_,x]) = SOME x
- | functions func_args = projection func_args
-
- fun relations ("<>",[x,y]) = SOME (x <> y)
- | relations _ = NONE
- in
- {functions = functions, relations = relations}
- end;
-
-fun fixedModulo {size = N} =
- let
- fun mod_N k = k mod N
-
- val one = mod_N 1
-
- fun mult (x,y) = mod_N (x * y)
-
- fun divides_N 0 = false
- | divides_N x = N mod x = 0
-
- val even_N = divides_N 2
-
- fun functions (numeral,[]) =
- Option.map mod_N (natFromString numeral handle Overflow => NONE)
- | functions ("suc",[x]) = SOME (if x = N - 1 then 0 else x + 1)
- | functions ("pre",[x]) = SOME (if x = 0 then N - 1 else x - 1)
- | functions ("~",[x]) = SOME (if x = 0 then 0 else N - x)
- | functions ("+",[x,y]) = SOME (mod_N (x + y))
- | functions ("-",[x,y]) = SOME (if x < y then N + x - y else x - y)
- | functions ("*",[x,y]) = SOME (mult (x,y))
- | functions ("exp",[x,y]) = SOME (exp mult x y one)
- | functions ("div",[x,y]) = if divides_N y then SOME (x div y) else NONE
- | functions ("mod",[x,y]) = if divides_N y then SOME (x mod y) else NONE
- | functions _ = NONE
-
- fun relations ("is_0",[x]) = SOME (x = 0)
- | relations ("divides",[x,y]) =
- if x = 0 then SOME (y = 0)
- else if divides_N x then SOME (y mod x = 0) else NONE
- | relations ("even",[x]) = if even_N then SOME (x mod 2 = 0) else NONE
- | relations ("odd",[x]) = if even_N then SOME (x mod 2 = 1) else NONE
- | relations _ = NONE
- in
- {functions = functions, relations = relations}
- end;
-
-local
- datatype onum = ONeg | ONum of int | OInf;
-
- val zero = ONum 0
- and one = ONum 1
- and two = ONum 2;
-
- fun suc (ONum x) = ONum (x + 1)
- | suc v = v;
-
- fun pre (ONum 0) = ONeg
- | pre (ONum x) = ONum (x - 1)
- | pre v = v;
-
- fun neg ONeg = NONE
- | neg (n as ONum 0) = SOME n
- | neg _ = SOME ONeg;
-
- fun add ONeg ONeg = SOME ONeg
- | add ONeg (ONum y) = if y = 0 then SOME ONeg else NONE
- | add ONeg OInf = NONE
- | add (ONum x) ONeg = if x = 0 then SOME ONeg else NONE
- | add (ONum x) (ONum y) = SOME (ONum (x + y))
- | add (ONum _) OInf = SOME OInf
- | add OInf ONeg = NONE
- | add OInf (ONum _) = SOME OInf
- | add OInf OInf = SOME OInf;
-
- fun sub ONeg ONeg = NONE
- | sub ONeg (ONum _) = SOME ONeg
- | sub ONeg OInf = SOME ONeg
- | sub (ONum _) ONeg = NONE
- | sub (ONum x) (ONum y) = SOME (if x < y then ONeg else ONum (x - y))
- | sub (ONum _) OInf = SOME ONeg
- | sub OInf ONeg = SOME OInf
- | sub OInf (ONum y) = if y = 0 then SOME OInf else NONE
- | sub OInf OInf = NONE;
-
- fun mult ONeg ONeg = NONE
- | mult ONeg (ONum y) = SOME (if y = 0 then zero else ONeg)
- | mult ONeg OInf = SOME ONeg
- | mult (ONum x) ONeg = SOME (if x = 0 then zero else ONeg)
- | mult (ONum x) (ONum y) = SOME (ONum (x * y))
- | mult (ONum x) OInf = SOME (if x = 0 then zero else OInf)
- | mult OInf ONeg = SOME ONeg
- | mult OInf (ONum y) = SOME (if y = 0 then zero else OInf)
- | mult OInf OInf = SOME OInf;
-
- fun exp ONeg ONeg = NONE
- | exp ONeg (ONum y) =
- if y = 0 then SOME one else if y mod 2 = 0 then NONE else SOME ONeg
- | exp ONeg OInf = NONE
- | exp (ONum x) ONeg = NONE
- | exp (ONum x) (ONum y) = SOME (ONum (intExp x y) handle Overflow => OInf)
- | exp (ONum x) OInf =
- SOME (if x = 0 then zero else if x = 1 then one else OInf)
- | exp OInf ONeg = NONE
- | exp OInf (ONum y) = SOME (if y = 0 then one else OInf)
- | exp OInf OInf = SOME OInf;
-
- fun odiv ONeg ONeg = NONE
- | odiv ONeg (ONum y) = if y = 1 then SOME ONeg else NONE
- | odiv ONeg OInf = NONE
- | odiv (ONum _) ONeg = NONE
- | odiv (ONum x) (ONum y) = if y = 0 then NONE else SOME (ONum (x div y))
- | odiv (ONum _) OInf = SOME zero
- | odiv OInf ONeg = NONE
- | odiv OInf (ONum y) = if y = 1 then SOME OInf else NONE
- | odiv OInf OInf = NONE;
-
- fun omod ONeg ONeg = NONE
- | omod ONeg (ONum y) = if y = 1 then SOME zero else NONE
- | omod ONeg OInf = NONE
- | omod (ONum _) ONeg = NONE
- | omod (ONum x) (ONum y) = if y = 0 then NONE else SOME (ONum (x mod y))
- | omod (x as ONum _) OInf = SOME x
- | omod OInf ONeg = NONE
- | omod OInf (ONum y) = if y = 1 then SOME OInf else NONE
- | omod OInf OInf = NONE;
-
- fun le ONeg ONeg = NONE
- | le ONeg (ONum y) = SOME true
- | le ONeg OInf = SOME true
- | le (ONum _) ONeg = SOME false
- | le (ONum x) (ONum y) = SOME (x <= y)
- | le (ONum _) OInf = SOME true
- | le OInf ONeg = SOME false
- | le OInf (ONum _) = SOME false
- | le OInf OInf = NONE;
-
- fun lt x y = Option.map not (le y x);
-
- fun ge x y = le y x;
-
- fun gt x y = lt y x;
-
- fun divides ONeg ONeg = NONE
- | divides ONeg (ONum y) = if y = 0 then SOME true else NONE
- | divides ONeg OInf = NONE
- | divides (ONum x) ONeg =
- if x = 0 then SOME false else if x = 1 then SOME true else NONE
- | divides (ONum x) (ONum y) = SOME (Useful.divides x y)
- | divides (ONum x) OInf =
- if x = 0 then SOME false else if x = 1 then SOME true else NONE
- | divides OInf ONeg = NONE
- | divides OInf (ONum y) = SOME (y = 0)
- | divides OInf OInf = NONE;
-
- fun even n = divides two n;
-
- fun odd n = Option.map not (even n);
-
- fun fixedOverflow mk_onum dest_onum =
- let
- fun partial_dest_onum NONE = NONE
- | partial_dest_onum (SOME n) = dest_onum n
-
- fun functions (numeral,[]) =
- (case (natFromString numeral handle Overflow => NONE) of
- NONE => NONE
- | SOME n => dest_onum (ONum n))
- | functions ("suc",[x]) = dest_onum (suc (mk_onum x))
- | functions ("pre",[x]) = dest_onum (pre (mk_onum x))
- | functions ("~",[x]) = partial_dest_onum (neg (mk_onum x))
- | functions ("+",[x,y]) =
- partial_dest_onum (add (mk_onum x) (mk_onum y))
- | functions ("-",[x,y]) =
- partial_dest_onum (sub (mk_onum x) (mk_onum y))
- | functions ("*",[x,y]) =
- partial_dest_onum (mult (mk_onum x) (mk_onum y))
- | functions ("exp",[x,y]) =
- partial_dest_onum (exp (mk_onum x) (mk_onum y))
- | functions ("div",[x,y]) =
- partial_dest_onum (odiv (mk_onum x) (mk_onum y))
- | functions ("mod",[x,y]) =
- partial_dest_onum (omod (mk_onum x) (mk_onum y))
- | functions _ = NONE
-
- fun relations ("is_0",[x]) = SOME (mk_onum x = zero)
- | relations ("<=",[x,y]) = le (mk_onum x) (mk_onum y)
- | relations ("<",[x,y]) = lt (mk_onum x) (mk_onum y)
- | relations (">=",[x,y]) = ge (mk_onum x) (mk_onum y)
- | relations (">",[x,y]) = gt (mk_onum x) (mk_onum y)
- | relations ("divides",[x,y]) = divides (mk_onum x) (mk_onum y)
- | relations ("even",[x]) = even (mk_onum x)
- | relations ("odd",[x]) = odd (mk_onum x)
- | relations _ = NONE
- in
- {functions = functions, relations = relations}
- end;
-in
- fun fixedOverflowNum {size = N} =
- let
- val oinf = N - 1
-
- fun mk_onum x = if x = oinf then OInf else ONum x
-
- fun dest_onum ONeg = NONE
- | dest_onum (ONum x) = SOME (if x < oinf then x else oinf)
- | dest_onum OInf = SOME oinf
- in
- fixedOverflow mk_onum dest_onum
- end;
-
- fun fixedOverflowInt {size = N} =
- let
- val oinf = N - 2
- val oneg = N - 1
-
- fun mk_onum x =
- if x = oneg then ONeg else if x = oinf then OInf else ONum x
-
- fun dest_onum ONeg = SOME oneg
- | dest_onum (ONum x) = SOME (if x < oinf then x else oinf)
- | dest_onum OInf = SOME oinf
- in
- fixedOverflow mk_onum dest_onum
- end;
-end;
-
-fun fixedSet {size = N} =
- let
- val M =
- let
- fun f 0 acc = acc
- | f x acc = f (x div 2) (acc + 1)
- in
- f N 0
- end
-
- val univ = IntSet.fromList (interval 0 M)
-
- val mk_set =
- let
- fun f _ s 0 = s
- | f k s x =
- let
- val s = if x mod 2 = 0 then s else IntSet.add s k
- in
- f (k + 1) s (x div 2)
- end
- in
- f 0 IntSet.empty
- end
-
- fun dest_set s =
- let
- fun f 0 x = x
- | f k x =
- let
- val k = k - 1
- in
- f k (if IntSet.member k s then 2 * x + 1 else 2 * x)
- end
-
- val x = case IntSet.findr (K true) s of NONE => 0 | SOME k => f k 1
- in
- if x < N then SOME x else NONE
- end
-
- fun functions ("empty",[]) = dest_set IntSet.empty
- | functions ("univ",[]) = dest_set univ
- | functions ("union",[x,y]) =
- dest_set (IntSet.union (mk_set x) (mk_set y))
- | functions ("intersect",[x,y]) =
- dest_set (IntSet.intersect (mk_set x) (mk_set y))
- | functions ("compl",[x]) =
- dest_set (IntSet.difference univ (mk_set x))
- | functions ("card",[x]) = SOME (IntSet.size (mk_set x))
- | functions _ = NONE
-
- fun relations ("in",[x,y]) = SOME (IntSet.member (x mod M) (mk_set y))
- | relations ("subset",[x,y]) =
- SOME (IntSet.subset (mk_set x) (mk_set y))
- | relations _ = NONE
- in
- {functions = functions, relations = relations}
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of random finite models. *)
-(* ------------------------------------------------------------------------- *)
-
-type parameters = {size : int, fixed : fixed};
-
-datatype model =
- Model of
- {size : int,
- fixed : fixedModel,
- functions : (Term.functionName * int list, int) Map.map ref,
- relations : (Atom.relationName * int list, bool) Map.map ref};
-
-local
- fun cmp ((n1,l1),(n2,l2)) =
- case String.compare (n1,n2) of
- LESS => LESS
- | EQUAL => lexCompare Int.compare (l1,l2)
- | GREATER => GREATER;
-in
- fun new {size = N, fixed} =
- Model
- {size = N,
- fixed = fixed {size = N},
- functions = ref (Map.new cmp),
- relations = ref (Map.new cmp)};
-end;
-
-fun size (Model {size = s, ...}) = s;
-
-(* ------------------------------------------------------------------------- *)
-(* Valuations. *)
-(* ------------------------------------------------------------------------- *)
-
-type valuation = int NameMap.map;
-
-val valuationEmpty : valuation = NameMap.new ();
-
-fun valuationRandom {size = N} vs =
- let
- fun f (v,V) = NameMap.insert V (v, Portable.randomInt N)
- in
- NameSet.foldl f valuationEmpty vs
- end;
-
-fun valuationFold {size = N} vs f =
- let
- val vs = NameSet.toList vs
-
- fun inc [] _ = NONE
- | inc (v :: l) V =
- case NameMap.peek V v of
- NONE => raise Bug "Model.valuationFold"
- | SOME k =>
- let
- val k = if k = N - 1 then 0 else k + 1
- val V = NameMap.insert V (v,k)
- in
- if k = 0 then inc l V else SOME V
- end
-
- val zero = foldl (fn (v,V) => NameMap.insert V (v,0)) valuationEmpty vs
-
- fun fold V acc =
- let
- val acc = f (V,acc)
- in
- case inc vs V of NONE => acc | SOME V => fold V acc
- end
- in
- fold zero
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Interpreting terms and formulas in the model. *)
-(* ------------------------------------------------------------------------- *)
-
-fun interpretTerm M V =
- let
- val Model {size = N, fixed, functions, ...} = M
- val {functions = fixed_functions, ...} = fixed
-
- fun interpret (Term.Var v) =
- (case NameMap.peek V v of
- NONE => raise Error "Model.interpretTerm: incomplete valuation"
- | SOME i => i)
- | interpret (tm as Term.Fn f_tms) =
- let
- val (f,tms) =
- case Term.stripComb tm of
- (_,[]) => f_tms
- | (v as Term.Var _, tms) => (".", v :: tms)
- | (Term.Fn (f,tms), tms') => (f, tms @ tms')
- val elts = map interpret tms
- val f_elts = (f,elts)
- val ref funcs = functions
- in
- case Map.peek funcs f_elts of
- SOME k => k
- | NONE =>
- let
- val k =
- case fixed_functions f_elts of
- SOME k => k
- | NONE => Portable.randomInt N
-
- val () = functions := Map.insert funcs (f_elts,k)
- in
- k
- end
- end;
- in
- interpret
- end;
-
-fun interpretAtom M V (r,tms) =
- let
- val Model {fixed,relations,...} = M
- val {relations = fixed_relations, ...} = fixed
-
- val elts = map (interpretTerm M V) tms
- val r_elts = (r,elts)
- val ref rels = relations
- in
- case Map.peek rels r_elts of
- SOME b => b
- | NONE =>
- let
- val b =
- case fixed_relations r_elts of
- SOME b => b
- | NONE => Portable.randomBool ()
-
- val () = relations := Map.insert rels (r_elts,b)
- in
- b
- end
- end;
-
-fun interpretFormula M =
- let
- val Model {size = N, ...} = M
-
- fun interpret _ Formula.True = true
- | interpret _ Formula.False = false
- | interpret V (Formula.Atom atm) = interpretAtom M V atm
- | interpret V (Formula.Not p) = not (interpret V p)
- | interpret V (Formula.Or (p,q)) = interpret V p orelse interpret V q
- | interpret V (Formula.And (p,q)) = interpret V p andalso interpret V q
- | interpret V (Formula.Imp (p,q)) =
- interpret V (Formula.Or (Formula.Not p, q))
- | interpret V (Formula.Iff (p,q)) = interpret V p = interpret V q
- | interpret V (Formula.Forall (v,p)) = interpret' V v p N
- | interpret V (Formula.Exists (v,p)) =
- interpret V (Formula.Not (Formula.Forall (v, Formula.Not p)))
- and interpret' _ _ _ 0 = true
- | interpret' V v p i =
- let
- val i = i - 1
- val V' = NameMap.insert V (v,i)
- in
- interpret V' p andalso interpret' V v p i
- end
- in
- interpret
- end;
-
-fun interpretLiteral M V (true,atm) = interpretAtom M V atm
- | interpretLiteral M V (false,atm) = not (interpretAtom M V atm);
-
-fun interpretClause M V cl = LiteralSet.exists (interpretLiteral M V) cl;
-
-(* ------------------------------------------------------------------------- *)
-(* Check whether random groundings of a formula are true in the model. *)
-(* Note: if it's cheaper, a systematic check will be performed instead. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun checkGen freeVars interpret {maxChecks} M x =
- let
- val Model {size = N, ...} = M
-
- fun score (V,{T,F}) =
- if interpret M V x then {T = T + 1, F = F} else {T = T, F = F + 1}
-
- val vs = freeVars x
-
- fun randomCheck acc = score (valuationRandom {size = N} vs, acc)
-
- val small =
- intExp N (NameSet.size vs) <= maxChecks handle Overflow => false
- in
- if small then valuationFold {size = N} vs score {T = 0, F = 0}
- else funpow maxChecks randomCheck {T = 0, F = 0}
- end;
-in
- val checkAtom = checkGen Atom.freeVars interpretAtom;
-
- val checkFormula = checkGen Formula.freeVars interpretFormula;
-
- val checkLiteral = checkGen Literal.freeVars interpretLiteral;
-
- val checkClause = checkGen LiteralSet.freeVars interpretClause;
-end;
-
-end
--- a/src/Tools/Metis/src/Name.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-(* ========================================================================= *)
-(* NAMES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Name =
-sig
-
-type name = string
-
-val compare : name * name -> order
-
-val pp : name Parser.pp
-
-end
--- a/src/Tools/Metis/src/Name.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,85 +0,0 @@
-(* ========================================================================= *)
-(* NAMES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Name :> Name =
-struct
-
-type name = string;
-
-val compare = String.compare;
-
-val pp = Parser.ppString;
-
-end
-
-structure NameOrdered =
-struct type t = Name.name val compare = Name.compare end
-
-structure NameSet =
-struct
-
- local
- structure S = ElementSet (NameOrdered);
- in
- open S;
- end;
-
- val pp =
- Parser.ppMap
- toList
- (Parser.ppBracket "{" "}" (Parser.ppSequence "," Name.pp));
-
-end
-
-structure NameMap = KeyMap (NameOrdered);
-
-structure NameArity =
-struct
-
-type nameArity = Name.name * int;
-
-fun name ((n,_) : nameArity) = n;
-
-fun arity ((_,i) : nameArity) = i;
-
-fun nary i n_i = arity n_i = i;
-
-val nullary = nary 0
-and unary = nary 1
-and binary = nary 2
-and ternary = nary 3;
-
-fun compare ((n1,i1),(n2,i2)) =
- case Name.compare (n1,n2) of
- LESS => LESS
- | EQUAL => Int.compare (i1,i2)
- | GREATER => GREATER;
-
-val pp = Parser.ppMap (fn (n,i) => n ^ "/" ^ Int.toString i) Parser.ppString;
-
-end
-
-structure NameArityOrdered =
-struct type t = NameArity.nameArity val compare = NameArity.compare end
-
-structure NameAritySet =
-struct
-
- local
- structure S = ElementSet (NameArityOrdered);
- in
- open S;
- end;
-
- val allNullary = all NameArity.nullary;
-
- val pp =
- Parser.ppMap
- toList
- (Parser.ppBracket "{" "}" (Parser.ppSequence "," NameArity.pp));
-
-end
-
-structure NameArityMap = KeyMap (NameArityOrdered);
--- a/src/Tools/Metis/src/Normalize.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,21 +0,0 @@
-(* ========================================================================= *)
-(* NORMALIZING FORMULAS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Normalize =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Negation normal form. *)
-(* ------------------------------------------------------------------------- *)
-
-val nnf : Formula.formula -> Formula.formula
-
-(* ------------------------------------------------------------------------- *)
-(* Conjunctive normal form. *)
-(* ------------------------------------------------------------------------- *)
-
-val cnf : Formula.formula -> Formula.formula list
-
-end
--- a/src/Tools/Metis/src/Normalize.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1062 +0,0 @@
-(* ========================================================================= *)
-(* NORMALIZING FORMULAS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Normalize :> Normalize =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Counting the clauses that would be generated by conjunctive normal form. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype count = Count of {positive : real, negative : real};
-
-fun positive (Count {positive = p, ...}) = p;
-
-fun negative (Count {negative = n, ...}) = n;
-
-fun countNegate (Count {positive = p, negative = n}) =
- Count {positive = n, negative = p};
-
-fun countEqualish count1 count2 =
- let
- val Count {positive = p1, negative = n1} = count1
- and Count {positive = p2, negative = n2} = count2
- in
- Real.abs (p1 - p2) < 0.5 andalso Real.abs (n1 - n2) < 0.5
- end;
-
-val countTrue = Count {positive = 0.0, negative = 1.0};
-
-val countFalse = Count {positive = 1.0, negative = 0.0};
-
-val countLiteral = Count {positive = 1.0, negative = 1.0};
-
-fun countAnd2 (count1,count2) =
- let
- val Count {positive = p1, negative = n1} = count1
- and Count {positive = p2, negative = n2} = count2
- val p = p1 + p2
- and n = n1 * n2
- in
- Count {positive = p, negative = n}
- end;
-
-fun countOr2 (count1,count2) =
- let
- val Count {positive = p1, negative = n1} = count1
- and Count {positive = p2, negative = n2} = count2
- val p = p1 * p2
- and n = n1 + n2
- in
- Count {positive = p, negative = n}
- end;
-
-(*** Is this associative? ***)
-fun countXor2 (count1,count2) =
- let
- val Count {positive = p1, negative = n1} = count1
- and Count {positive = p2, negative = n2} = count2
- val p = p1 * p2 + n1 * n2
- and n = p1 * n2 + n1 * p2
- in
- Count {positive = p, negative = n}
- end;
-
-fun countDefinition body_count = countXor2 (countLiteral,body_count);
-
-(* ------------------------------------------------------------------------- *)
-(* A type of normalized formula. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formula =
- True
- | False
- | Literal of NameSet.set * Literal.literal
- | And of NameSet.set * count * formula Set.set
- | Or of NameSet.set * count * formula Set.set
- | Xor of NameSet.set * count * bool * formula Set.set
- | Exists of NameSet.set * count * NameSet.set * formula
- | Forall of NameSet.set * count * NameSet.set * formula;
-
-fun compare f1_f2 =
- case f1_f2 of
- (True,True) => EQUAL
- | (True,_) => LESS
- | (_,True) => GREATER
- | (False,False) => EQUAL
- | (False,_) => LESS
- | (_,False) => GREATER
- | (Literal (_,l1), Literal (_,l2)) => Literal.compare (l1,l2)
- | (Literal _, _) => LESS
- | (_, Literal _) => GREATER
- | (And (_,_,s1), And (_,_,s2)) => Set.compare (s1,s2)
- | (And _, _) => LESS
- | (_, And _) => GREATER
- | (Or (_,_,s1), Or (_,_,s2)) => Set.compare (s1,s2)
- | (Or _, _) => LESS
- | (_, Or _) => GREATER
- | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) =>
- (case boolCompare (p1,p2) of
- LESS => LESS
- | EQUAL => Set.compare (s1,s2)
- | GREATER => GREATER)
- | (Xor _, _) => LESS
- | (_, Xor _) => GREATER
- | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) =>
- (case NameSet.compare (n1,n2) of
- LESS => LESS
- | EQUAL => compare (f1,f2)
- | GREATER => GREATER)
- | (Exists _, _) => LESS
- | (_, Exists _) => GREATER
- | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) =>
- (case NameSet.compare (n1,n2) of
- LESS => LESS
- | EQUAL => compare (f1,f2)
- | GREATER => GREATER);
-
-val empty = Set.empty compare;
-
-val singleton = Set.singleton compare;
-
-local
- fun neg True = False
- | neg False = True
- | neg (Literal (fv,lit)) = Literal (fv, Literal.negate lit)
- | neg (And (fv,c,s)) = Or (fv, countNegate c, neg_set s)
- | neg (Or (fv,c,s)) = And (fv, countNegate c, neg_set s)
- | neg (Xor (fv,c,p,s)) = Xor (fv, c, not p, s)
- | neg (Exists (fv,c,n,f)) = Forall (fv, countNegate c, n, neg f)
- | neg (Forall (fv,c,n,f)) = Exists (fv, countNegate c, n, neg f)
-
- and neg_set s = Set.foldl neg_elt empty s
-
- and neg_elt (f,s) = Set.add s (neg f);
-in
- val negate = neg;
-
- val negateSet = neg_set;
-end;
-
-fun negateMember x s = Set.member (negate x) s;
-
-local
- fun member s x = negateMember x s;
-in
- fun negateDisjoint s1 s2 =
- if Set.size s1 < Set.size s2 then not (Set.exists (member s2) s1)
- else not (Set.exists (member s1) s2);
-end;
-
-fun polarity True = true
- | polarity False = false
- | polarity (Literal (_,(pol,_))) = not pol
- | polarity (And _) = true
- | polarity (Or _) = false
- | polarity (Xor (_,_,pol,_)) = pol
- | polarity (Exists _) = true
- | polarity (Forall _) = false;
-
-(*DEBUG
-val polarity = fn f =>
- let
- val res1 = compare (f, negate f) = LESS
- val res2 = polarity f
- val _ = res1 = res2 orelse raise Bug "polarity"
- in
- res2
- end;
-*)
-
-fun applyPolarity true fm = fm
- | applyPolarity false fm = negate fm;
-
-fun freeVars True = NameSet.empty
- | freeVars False = NameSet.empty
- | freeVars (Literal (fv,_)) = fv
- | freeVars (And (fv,_,_)) = fv
- | freeVars (Or (fv,_,_)) = fv
- | freeVars (Xor (fv,_,_,_)) = fv
- | freeVars (Exists (fv,_,_,_)) = fv
- | freeVars (Forall (fv,_,_,_)) = fv;
-
-fun freeIn v fm = NameSet.member v (freeVars fm);
-
-val freeVarsSet =
- let
- fun free (fm,acc) = NameSet.union (freeVars fm) acc
- in
- Set.foldl free NameSet.empty
- end;
-
-fun count True = countTrue
- | count False = countFalse
- | count (Literal _) = countLiteral
- | count (And (_,c,_)) = c
- | count (Or (_,c,_)) = c
- | count (Xor (_,c,p,_)) = if p then c else countNegate c
- | count (Exists (_,c,_,_)) = c
- | count (Forall (_,c,_,_)) = c;
-
-val countAndSet =
- let
- fun countAnd (fm,c) = countAnd2 (count fm, c)
- in
- Set.foldl countAnd countTrue
- end;
-
-val countOrSet =
- let
- fun countOr (fm,c) = countOr2 (count fm, c)
- in
- Set.foldl countOr countFalse
- end;
-
-val countXorSet =
- let
- fun countXor (fm,c) = countXor2 (count fm, c)
- in
- Set.foldl countXor countFalse
- end;
-
-fun And2 (False,_) = False
- | And2 (_,False) = False
- | And2 (True,f2) = f2
- | And2 (f1,True) = f1
- | And2 (f1,f2) =
- let
- val (fv1,c1,s1) =
- case f1 of
- And fv_c_s => fv_c_s
- | _ => (freeVars f1, count f1, singleton f1)
-
- and (fv2,c2,s2) =
- case f2 of
- And fv_c_s => fv_c_s
- | _ => (freeVars f2, count f2, singleton f2)
- in
- if not (negateDisjoint s1 s2) then False
- else
- let
- val s = Set.union s1 s2
- in
- case Set.size s of
- 0 => True
- | 1 => Set.pick s
- | n =>
- if n = Set.size s1 + Set.size s2 then
- And (NameSet.union fv1 fv2, countAnd2 (c1,c2), s)
- else
- And (freeVarsSet s, countAndSet s, s)
- end
- end;
-
-val AndList = foldl And2 True;
-
-val AndSet = Set.foldl And2 True;
-
-fun Or2 (True,_) = True
- | Or2 (_,True) = True
- | Or2 (False,f2) = f2
- | Or2 (f1,False) = f1
- | Or2 (f1,f2) =
- let
- val (fv1,c1,s1) =
- case f1 of
- Or fv_c_s => fv_c_s
- | _ => (freeVars f1, count f1, singleton f1)
-
- and (fv2,c2,s2) =
- case f2 of
- Or fv_c_s => fv_c_s
- | _ => (freeVars f2, count f2, singleton f2)
- in
- if not (negateDisjoint s1 s2) then True
- else
- let
- val s = Set.union s1 s2
- in
- case Set.size s of
- 0 => False
- | 1 => Set.pick s
- | n =>
- if n = Set.size s1 + Set.size s2 then
- Or (NameSet.union fv1 fv2, countOr2 (c1,c2), s)
- else
- Or (freeVarsSet s, countOrSet s, s)
- end
- end;
-
-val OrList = foldl Or2 False;
-
-val OrSet = Set.foldl Or2 False;
-
-fun pushOr2 (f1,f2) =
- let
- val s1 = case f1 of And (_,_,s) => s | _ => singleton f1
- and s2 = case f2 of And (_,_,s) => s | _ => singleton f2
-
- fun g x1 (x2,acc) = And2 (Or2 (x1,x2), acc)
- fun f (x1,acc) = Set.foldl (g x1) acc s2
- in
- Set.foldl f True s1
- end;
-
-val pushOrList = foldl pushOr2 False;
-
-local
- fun normalize fm =
- let
- val p = polarity fm
- val fm = applyPolarity p fm
- in
- (freeVars fm, count fm, p, singleton fm)
- end;
-in
- fun Xor2 (False,f2) = f2
- | Xor2 (f1,False) = f1
- | Xor2 (True,f2) = negate f2
- | Xor2 (f1,True) = negate f1
- | Xor2 (f1,f2) =
- let
- val (fv1,c1,p1,s1) = case f1 of Xor x => x | _ => normalize f1
- and (fv2,c2,p2,s2) = case f2 of Xor x => x | _ => normalize f2
-
- val s = Set.symmetricDifference s1 s2
-
- val fm =
- case Set.size s of
- 0 => False
- | 1 => Set.pick s
- | n =>
- if n = Set.size s1 + Set.size s2 then
- Xor (NameSet.union fv1 fv2, countXor2 (c1,c2), true, s)
- else
- Xor (freeVarsSet s, countXorSet s, true, s)
-
- val p = p1 = p2
- in
- applyPolarity p fm
- end;
-end;
-
-val XorList = foldl Xor2 False;
-
-val XorSet = Set.foldl Xor2 False;
-
-fun XorPolarityList (p,l) = applyPolarity p (XorList l);
-
-fun XorPolaritySet (p,s) = applyPolarity p (XorSet s);
-
-fun destXor (Xor (_,_,p,s)) =
- let
- val (fm1,s) = Set.deletePick s
- val fm2 =
- if Set.size s = 1 then applyPolarity p (Set.pick s)
- else Xor (freeVarsSet s, countXorSet s, p, s)
- in
- (fm1,fm2)
- end
- | destXor _ = raise Error "destXor";
-
-fun pushXor fm =
- let
- val (f1,f2) = destXor fm
- val f1' = negate f1
- and f2' = negate f2
- in
- And2 (Or2 (f1,f2), Or2 (f1',f2'))
- end;
-
-fun Exists1 (v,init_fm) =
- let
- fun exists_gen fm =
- let
- val fv = NameSet.delete (freeVars fm) v
- val c = count fm
- val n = NameSet.singleton v
- in
- Exists (fv,c,n,fm)
- end
-
- fun exists fm = if freeIn v fm then exists_free fm else fm
-
- and exists_free (Or (_,_,s)) = OrList (Set.transform exists s)
- | exists_free (fm as And (_,_,s)) =
- let
- val sv = Set.filter (freeIn v) s
- in
- if Set.size sv <> 1 then exists_gen fm
- else
- let
- val fm = Set.pick sv
- val s = Set.delete s fm
- in
- And2 (exists_free fm, AndSet s)
- end
- end
- | exists_free (Exists (fv,c,n,f)) =
- Exists (NameSet.delete fv v, c, NameSet.add n v, f)
- | exists_free fm = exists_gen fm
- in
- exists init_fm
- end;
-
-fun ExistsList (vs,f) = foldl Exists1 f vs;
-
-fun ExistsSet (n,f) = NameSet.foldl Exists1 f n;
-
-fun Forall1 (v,init_fm) =
- let
- fun forall_gen fm =
- let
- val fv = NameSet.delete (freeVars fm) v
- val c = count fm
- val n = NameSet.singleton v
- in
- Forall (fv,c,n,fm)
- end
-
- fun forall fm = if freeIn v fm then forall_free fm else fm
-
- and forall_free (And (_,_,s)) = AndList (Set.transform forall s)
- | forall_free (fm as Or (_,_,s)) =
- let
- val sv = Set.filter (freeIn v) s
- in
- if Set.size sv <> 1 then forall_gen fm
- else
- let
- val fm = Set.pick sv
- val s = Set.delete s fm
- in
- Or2 (forall_free fm, OrSet s)
- end
- end
- | forall_free (Forall (fv,c,n,f)) =
- Forall (NameSet.delete fv v, c, NameSet.add n v, f)
- | forall_free fm = forall_gen fm
- in
- forall init_fm
- end;
-
-fun ForallList (vs,f) = foldl Forall1 f vs;
-
-fun ForallSet (n,f) = NameSet.foldl Forall1 f n;
-
-local
- fun subst_fv fvSub =
- let
- fun add_fv (v,s) = NameSet.union (NameMap.get fvSub v) s
- in
- NameSet.foldl add_fv NameSet.empty
- end;
-
- fun subst_rename (v,(avoid,bv,sub,domain,fvSub)) =
- let
- val v' = Term.variantPrime avoid v
- val avoid = NameSet.add avoid v'
- val bv = NameSet.add bv v'
- val sub = Subst.insert sub (v, Term.Var v')
- val domain = NameSet.add domain v
- val fvSub = NameMap.insert fvSub (v, NameSet.singleton v')
- in
- (avoid,bv,sub,domain,fvSub)
- end;
-
- fun subst_check sub domain fvSub fm =
- let
- val domain = NameSet.intersect domain (freeVars fm)
- in
- if NameSet.null domain then fm
- else subst_domain sub domain fvSub fm
- end
-
- and subst_domain sub domain fvSub fm =
- case fm of
- Literal (fv,lit) =>
- let
- val fv = NameSet.difference fv domain
- val fv = NameSet.union fv (subst_fv fvSub domain)
- val lit = Literal.subst sub lit
- in
- Literal (fv,lit)
- end
- | And (_,_,s) =>
- AndList (Set.transform (subst_check sub domain fvSub) s)
- | Or (_,_,s) =>
- OrList (Set.transform (subst_check sub domain fvSub) s)
- | Xor (_,_,p,s) =>
- XorPolarityList (p, Set.transform (subst_check sub domain fvSub) s)
- | Exists fv_c_n_f => subst_quant Exists sub domain fvSub fv_c_n_f
- | Forall fv_c_n_f => subst_quant Forall sub domain fvSub fv_c_n_f
- | _ => raise Bug "subst_domain"
-
- and subst_quant quant sub domain fvSub (fv,c,bv,fm) =
- let
- val sub_fv = subst_fv fvSub domain
- val fv = NameSet.union sub_fv (NameSet.difference fv domain)
- val captured = NameSet.intersect bv sub_fv
- val bv = NameSet.difference bv captured
- val avoid = NameSet.union fv bv
- val (_,bv,sub,domain,fvSub) =
- NameSet.foldl subst_rename (avoid,bv,sub,domain,fvSub) captured
- val fm = subst_domain sub domain fvSub fm
- in
- quant (fv,c,bv,fm)
- end;
-in
- fun subst sub =
- let
- fun mk_dom (v,tm,(d,fv)) =
- (NameSet.add d v, NameMap.insert fv (v, Term.freeVars tm))
-
- val domain_fvSub = (NameSet.empty, NameMap.new ())
- val (domain,fvSub) = Subst.foldl mk_dom domain_fvSub sub
- in
- subst_check sub domain fvSub
- end;
-end;
-
-fun fromFormula fm =
- case fm of
- Formula.True => True
- | Formula.False => False
- | Formula.Atom atm => Literal (Atom.freeVars atm, (true,atm))
- | Formula.Not p => negateFromFormula p
- | Formula.And (p,q) => And2 (fromFormula p, fromFormula q)
- | Formula.Or (p,q) => Or2 (fromFormula p, fromFormula q)
- | Formula.Imp (p,q) => Or2 (negateFromFormula p, fromFormula q)
- | Formula.Iff (p,q) => Xor2 (negateFromFormula p, fromFormula q)
- | Formula.Forall (v,p) => Forall1 (v, fromFormula p)
- | Formula.Exists (v,p) => Exists1 (v, fromFormula p)
-
-and negateFromFormula fm =
- case fm of
- Formula.True => False
- | Formula.False => True
- | Formula.Atom atm => Literal (Atom.freeVars atm, (false,atm))
- | Formula.Not p => fromFormula p
- | Formula.And (p,q) => Or2 (negateFromFormula p, negateFromFormula q)
- | Formula.Or (p,q) => And2 (negateFromFormula p, negateFromFormula q)
- | Formula.Imp (p,q) => And2 (fromFormula p, negateFromFormula q)
- | Formula.Iff (p,q) => Xor2 (fromFormula p, fromFormula q)
- | Formula.Forall (v,p) => Exists1 (v, negateFromFormula p)
- | Formula.Exists (v,p) => Forall1 (v, negateFromFormula p);
-
-local
- fun lastElt (s : formula Set.set) =
- case Set.findr (K true) s of
- NONE => raise Bug "lastElt: empty set"
- | SOME fm => fm;
-
- fun negateLastElt s =
- let
- val fm = lastElt s
- in
- Set.add (Set.delete s fm) (negate fm)
- end;
-
- fun form fm =
- case fm of
- True => Formula.True
- | False => Formula.False
- | Literal (_,lit) => Literal.toFormula lit
- | And (_,_,s) => Formula.listMkConj (Set.transform form s)
- | Or (_,_,s) => Formula.listMkDisj (Set.transform form s)
- | Xor (_,_,p,s) =>
- let
- val s = if p then negateLastElt s else s
- in
- Formula.listMkEquiv (Set.transform form s)
- end
- | Exists (_,_,n,f) => Formula.listMkExists (NameSet.toList n, form f)
- | Forall (_,_,n,f) => Formula.listMkForall (NameSet.toList n, form f);
-in
- val toFormula = form;
-end;
-
-val pp = Parser.ppMap toFormula Formula.pp;
-
-val toString = Parser.toString pp;
-
-(* ------------------------------------------------------------------------- *)
-(* Negation normal form. *)
-(* ------------------------------------------------------------------------- *)
-
-fun nnf fm = toFormula (fromFormula fm);
-
-(* ------------------------------------------------------------------------- *)
-(* Simplifying with definitions. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype simplify =
- Simplify of
- {formula : (formula,formula) Map.map,
- andSet : (formula Set.set * formula) list,
- orSet : (formula Set.set * formula) list,
- xorSet : (formula Set.set * formula) list};
-
-val simplifyEmpty =
- Simplify {formula = Map.new compare, andSet = [], orSet = [], xorSet = []};
-
-local
- fun simpler fm s =
- Set.size s <> 1 orelse
- case Set.pick s of
- True => false
- | False => false
- | Literal _ => false
- | _ => true;
-
- fun addSet set_defs body_def =
- let
- fun def_body_size (body,_) = Set.size body
-
- val body_size = def_body_size body_def
-
- val (body,_) = body_def
-
- fun add acc [] = List.revAppend (acc,[body_def])
- | add acc (l as (bd as (b,_)) :: bds) =
- case Int.compare (def_body_size bd, body_size) of
- LESS => List.revAppend (acc, body_def :: l)
- | EQUAL => if Set.equal b body then List.revAppend (acc,l)
- else add (bd :: acc) bds
- | GREATER => add (bd :: acc) bds
- in
- add [] set_defs
- end;
-
- fun add simp (body,False) = add simp (negate body, True)
- | add simp (True,_) = simp
- | add (Simplify {formula,andSet,orSet,xorSet}) (And (_,_,s), def) =
- let
- val andSet = addSet andSet (s,def)
- and orSet = addSet orSet (negateSet s, negate def)
- in
- Simplify
- {formula = formula,
- andSet = andSet, orSet = orSet, xorSet = xorSet}
- end
- | add (Simplify {formula,andSet,orSet,xorSet}) (Or (_,_,s), def) =
- let
- val orSet = addSet orSet (s,def)
- and andSet = addSet andSet (negateSet s, negate def)
- in
- Simplify
- {formula = formula,
- andSet = andSet, orSet = orSet, xorSet = xorSet}
- end
- | add simp (Xor (_,_,p,s), def) =
- let
- val simp = addXorSet simp (s, applyPolarity p def)
- in
- case def of
- True =>
- let
- fun addXorLiteral (fm as Literal _, simp) =
- let
- val s = Set.delete s fm
- in
- if not (simpler fm s) then simp
- else addXorSet simp (s, applyPolarity (not p) fm)
- end
- | addXorLiteral (_,simp) = simp
- in
- Set.foldl addXorLiteral simp s
- end
- | _ => simp
- end
- | add (simp as Simplify {formula,andSet,orSet,xorSet}) (body,def) =
- if Map.inDomain body formula then simp
- else
- let
- val formula = Map.insert formula (body,def)
- val formula = Map.insert formula (negate body, negate def)
- in
- Simplify
- {formula = formula,
- andSet = andSet, orSet = orSet, xorSet = xorSet}
- end
-
- and addXorSet (simp as Simplify {formula,andSet,orSet,xorSet}) (s,def) =
- if Set.size s = 1 then add simp (Set.pick s, def)
- else
- let
- val xorSet = addSet xorSet (s,def)
- in
- Simplify
- {formula = formula,
- andSet = andSet, orSet = orSet, xorSet = xorSet}
- end;
-in
- fun simplifyAdd simp fm = add simp (fm,True);
-end;
-
-local
- fun simplifySet set_defs set =
- let
- fun pred (s,_) = Set.subset s set
- in
- case List.find pred set_defs of
- NONE => NONE
- | SOME (s,f) => SOME (Set.add (Set.difference set s) f)
- end;
-in
- fun simplify (Simplify {formula,andSet,orSet,xorSet}) =
- let
- fun simp fm = simp_top (simp_sub fm)
-
- and simp_top (changed_fm as (_, And (_,_,s))) =
- (case simplifySet andSet s of
- NONE => changed_fm
- | SOME s => simp_top (true, AndSet s))
- | simp_top (changed_fm as (_, Or (_,_,s))) =
- (case simplifySet orSet s of
- NONE => changed_fm
- | SOME s => simp_top (true, OrSet s))
- | simp_top (changed_fm as (_, Xor (_,_,p,s))) =
- (case simplifySet xorSet s of
- NONE => changed_fm
- | SOME s => simp_top (true, XorPolaritySet (p,s)))
- | simp_top (changed_fm as (_,fm)) =
- (case Map.peek formula fm of
- NONE => changed_fm
- | SOME fm => simp_top (true,fm))
-
- and simp_sub fm =
- case fm of
- And (_,_,s) =>
- let
- val l = Set.transform simp s
- val changed = List.exists fst l
- val fm = if changed then AndList (map snd l) else fm
- in
- (changed,fm)
- end
- | Or (_,_,s) =>
- let
- val l = Set.transform simp s
- val changed = List.exists fst l
- val fm = if changed then OrList (map snd l) else fm
- in
- (changed,fm)
- end
- | Xor (_,_,p,s) =>
- let
- val l = Set.transform simp s
- val changed = List.exists fst l
- val fm = if changed then XorPolarityList (p, map snd l) else fm
- in
- (changed,fm)
- end
- | Exists (_,_,n,f) =>
- let
- val (changed,f) = simp f
- val fm = if changed then ExistsSet (n,f) else fm
- in
- (changed,fm)
- end
- | Forall (_,_,n,f) =>
- let
- val (changed,f) = simp f
- val fm = if changed then ForallSet (n,f) else fm
- in
- (changed,fm)
- end
- | _ => (false,fm);
- in
- fn fm => snd (simp fm)
- end;
-end;
-
-(*TRACE2
-val simplify = fn simp => fn fm =>
- let
- val fm' = simplify simp fm
- val () = if compare (fm,fm') = EQUAL then ()
- else (Parser.ppTrace pp "Normalize.simplify: fm" fm;
- Parser.ppTrace pp "Normalize.simplify: fm'" fm')
- in
- fm'
- end;
-*)
-
-(* ------------------------------------------------------------------------- *)
-(* Basic conjunctive normal form. *)
-(* ------------------------------------------------------------------------- *)
-
-val newSkolemFunction =
- let
- val counter : int NameMap.map ref = ref (NameMap.new ())
- in
- fn n => CRITICAL (fn () =>
- let
- val ref m = counter
- val i = Option.getOpt (NameMap.peek m n, 0)
- val () = counter := NameMap.insert m (n, i + 1)
- in
- "skolem_" ^ n ^ (if i = 0 then "" else "_" ^ Int.toString i)
- end)
- end;
-
-fun skolemize fv bv fm =
- let
- val fv = NameSet.transform Term.Var fv
-
- fun mk (v,s) = Subst.insert s (v, Term.Fn (newSkolemFunction v, fv))
- in
- subst (NameSet.foldl mk Subst.empty bv) fm
- end;
-
-local
- fun rename avoid fv bv fm =
- let
- val captured = NameSet.intersect avoid bv
- in
- if NameSet.null captured then fm
- else
- let
- fun ren (v,(a,s)) =
- let
- val v' = Term.variantPrime a v
- in
- (NameSet.add a v', Subst.insert s (v, Term.Var v'))
- end
-
- val avoid = NameSet.union (NameSet.union avoid fv) bv
-
- val (_,sub) = NameSet.foldl ren (avoid,Subst.empty) captured
- in
- subst sub fm
- end
- end;
-
- fun cnf avoid fm =
-(*TRACE5
- let
- val fm' = cnf' avoid fm
- val () = Parser.ppTrace pp "Normalize.cnf: fm" fm
- val () = Parser.ppTrace pp "Normalize.cnf: fm'" fm'
- in
- fm'
- end
- and cnf' avoid fm =
-*)
- case fm of
- True => True
- | False => False
- | Literal _ => fm
- | And (_,_,s) => AndList (Set.transform (cnf avoid) s)
- | Or (_,_,s) => pushOrList (snd (Set.foldl cnfOr (avoid,[]) s))
- | Xor _ => cnf avoid (pushXor fm)
- | Exists (fv,_,n,f) => cnf avoid (skolemize fv n f)
- | Forall (fv,_,n,f) => cnf avoid (rename avoid fv n f)
-
- and cnfOr (fm,(avoid,acc)) =
- let
- val fm = cnf avoid fm
- in
- (NameSet.union (freeVars fm) avoid, fm :: acc)
- end;
-in
- val basicCnf = cnf NameSet.empty;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Finding the formula definition that minimizes the number of clauses. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- type best = real * formula option;
-
- fun minBreak countClauses fm best =
- case fm of
- True => best
- | False => best
- | Literal _ => best
- | And (_,_,s) =>
- minBreakSet countClauses countAnd2 countTrue AndSet s best
- | Or (_,_,s) =>
- minBreakSet countClauses countOr2 countFalse OrSet s best
- | Xor (_,_,_,s) =>
- minBreakSet countClauses countXor2 countFalse XorSet s best
- | Exists (_,_,_,f) => minBreak countClauses f best
- | Forall (_,_,_,f) => minBreak countClauses f best
-
- and minBreakSet countClauses count2 count0 mkSet fmSet best =
- let
- fun cumulatives fms =
- let
- fun fwd (fm,(c1,s1,l)) =
- let
- val c1' = count2 (count fm, c1)
- and s1' = Set.add s1 fm
- in
- (c1', s1', (c1,s1,fm) :: l)
- end
-
- fun bwd ((c1,s1,fm),(c2,s2,l)) =
- let
- val c2' = count2 (count fm, c2)
- and s2' = Set.add s2 fm
- in
- (c2', s2', (c1,s1,fm,c2,s2) :: l)
- end
-
- val (c1,_,fms) = foldl fwd (count0,empty,[]) fms
- val (c2,_,fms) = foldl bwd (count0,empty,[]) fms
-
- val _ = countEqualish c1 c2 orelse raise Bug "cumulativeCounts"
- in
- fms
- end
-
- fun breakSing ((c1,_,fm,c2,_),best) =
- let
- val cFms = count2 (c1,c2)
- fun countCls cFm = countClauses (count2 (cFms,cFm))
- in
- minBreak countCls fm best
- end
-
- val breakSet1 =
- let
- fun break c1 s1 fm c2 (best as (bcl,_)) =
- if Set.null s1 then best
- else
- let
- val cDef = countDefinition (count2 (c1, count fm))
- val cFm = count2 (countLiteral,c2)
- val cl = positive cDef + countClauses cFm
- val better = cl < bcl - 0.5
- in
- if better then (cl, SOME (mkSet (Set.add s1 fm)))
- else best
- end
- in
- fn ((c1,s1,fm,c2,s2),best) =>
- break c1 s1 fm c2 (break c2 s2 fm c1 best)
- end
-
- val fms = Set.toList fmSet
-
- fun breakSet measure best =
- let
- val fms = sortMap (measure o count) Real.compare fms
- in
- foldl breakSet1 best (cumulatives fms)
- end
-
- val best = foldl breakSing best (cumulatives fms)
- val best = breakSet positive best
- val best = breakSet negative best
- val best = breakSet countClauses best
- in
- best
- end
-in
- fun minimumDefinition fm =
- let
- val countClauses = positive
- val cl = countClauses (count fm)
- in
- if cl < 1.5 then NONE
- else
- let
- val (cl',def) = minBreak countClauses fm (cl,NONE)
-(*TRACE1
- val () =
- case def of
- NONE => ()
- | SOME d =>
- Parser.ppTrace pp ("defCNF: before = " ^ Real.toString cl ^
- ", after = " ^ Real.toString cl' ^
- ", definition") d
-*)
- in
- def
- end
- end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Conjunctive normal form. *)
-(* ------------------------------------------------------------------------- *)
-
-val newDefinitionRelation =
- let
- val counter : int ref = ref 0
- in
- fn () => CRITICAL (fn () =>
- let
- val ref i = counter
- val () = counter := i + 1
- in
- "defCNF_" ^ Int.toString i
- end)
- end;
-
-fun newDefinition def =
- let
- val fv = freeVars def
- val atm = (newDefinitionRelation (), NameSet.transform Term.Var fv)
- val lit = Literal (fv,(true,atm))
- in
- Xor2 (lit,def)
- end;
-
-local
- fun def_cnf acc [] = acc
- | def_cnf acc ((prob,simp,fms) :: probs) =
- def_cnf_problem acc prob simp fms probs
-
- and def_cnf_problem acc prob _ [] probs = def_cnf (prob :: acc) probs
- | def_cnf_problem acc prob simp (fm :: fms) probs =
- def_cnf_formula acc prob simp (simplify simp fm) fms probs
-
- and def_cnf_formula acc prob simp fm fms probs =
- case fm of
- True => def_cnf_problem acc prob simp fms probs
- | False => def_cnf acc probs
- | And (_,_,s) => def_cnf_problem acc prob simp (Set.toList s @ fms) probs
- | Exists (fv,_,n,f) =>
- def_cnf_formula acc prob simp (skolemize fv n f) fms probs
- | Forall (_,_,_,f) => def_cnf_formula acc prob simp f fms probs
- | _ =>
- case minimumDefinition fm of
- NONE =>
- let
- val prob = fm :: prob
- and simp = simplifyAdd simp fm
- in
- def_cnf_problem acc prob simp fms probs
- end
- | SOME def =>
- let
- val def_fm = newDefinition def
- and fms = fm :: fms
- in
- def_cnf_formula acc prob simp def_fm fms probs
- end;
-
- fun cnf_prob prob = toFormula (AndList (map basicCnf prob));
-in
- fun cnf fm =
- let
- val fm = fromFormula fm
-(*TRACE2
- val () = Parser.ppTrace pp "Normalize.cnf: fm" fm
-*)
- val probs = def_cnf [] [([],simplifyEmpty,[fm])]
- in
- map cnf_prob probs
- end;
-end;
-
-end
--- a/src/Tools/Metis/src/Options.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,93 +0,0 @@
-(* ========================================================================= *)
-(* PROCESSING COMMAND LINE OPTIONS *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Options =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Option processors take an option with its associated arguments. *)
-(* ------------------------------------------------------------------------- *)
-
-type proc = string * string list -> unit
-
-type ('a,'x) mkProc = ('x -> proc) -> ('a -> 'x) -> proc
-
-(* ------------------------------------------------------------------------- *)
-(* One command line option: names, arguments, description and a processor. *)
-(* ------------------------------------------------------------------------- *)
-
-type opt =
- {switches : string list, arguments : string list,
- description : string, processor : proc}
-
-(* ------------------------------------------------------------------------- *)
-(* Option processors may raise an OptionExit exception. *)
-(* ------------------------------------------------------------------------- *)
-
-type optionExit = {message : string option, usage : bool, success : bool}
-
-exception OptionExit of optionExit
-
-(* ------------------------------------------------------------------------- *)
-(* Constructing option processors. *)
-(* ------------------------------------------------------------------------- *)
-
-val beginOpt : (string,'x) mkProc
-
-val endOpt : unit -> proc
-
-val stringOpt : (string,'x) mkProc
-
-val intOpt : int option * int option -> (int,'x) mkProc
-
-val realOpt : real option * real option -> (real,'x) mkProc
-
-val enumOpt : string list -> (string,'x) mkProc
-
-val optionOpt : string * ('a,'x) mkProc -> ('a option,'x) mkProc
-
-(* ------------------------------------------------------------------------- *)
-(* Basic options useful for all programs. *)
-(* ------------------------------------------------------------------------- *)
-
-val basicOptions : opt list
-
-(* ------------------------------------------------------------------------- *)
-(* All the command line options of a program. *)
-(* ------------------------------------------------------------------------- *)
-
-type allOptions =
- {name : string, version : string, header : string,
- footer : string, options : opt list}
-
-(* ------------------------------------------------------------------------- *)
-(* Usage information. *)
-(* ------------------------------------------------------------------------- *)
-
-val versionInformation : allOptions -> string
-
-val usageInformation : allOptions -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Exit the program gracefully. *)
-(* ------------------------------------------------------------------------- *)
-
-val exit : allOptions -> optionExit -> 'exit
-
-val succeed : allOptions -> 'exit
-
-val fail : allOptions -> string -> 'exit
-
-val usage : allOptions -> string -> 'exit
-
-val version : allOptions -> 'exit
-
-(* ------------------------------------------------------------------------- *)
-(* Process the command line options passed to the program. *)
-(* ------------------------------------------------------------------------- *)
-
-val processOptions : allOptions -> string list -> string list * string list
-
-end
--- a/src/Tools/Metis/src/Options.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,241 +0,0 @@
-(* ========================================================================= *)
-(* PROCESSING COMMAND LINE OPTIONS *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Options :> Options =
-struct
-
-infix ##
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* One command line option: names, arguments, description and a processor *)
-(* ------------------------------------------------------------------------- *)
-
-type proc = string * string list -> unit;
-
-type ('a,'x) mkProc = ('x -> proc) -> ('a -> 'x) -> proc;
-
-type opt = {switches : string list, arguments : string list,
- description : string, processor : proc};
-
-(* ------------------------------------------------------------------------- *)
-(* Option processors may raise an OptionExit exception *)
-(* ------------------------------------------------------------------------- *)
-
-type optionExit = {message : string option, usage : bool, success : bool};
-
-exception OptionExit of optionExit;
-
-(* ------------------------------------------------------------------------- *)
-(* Wrappers for option processors *)
-(* ------------------------------------------------------------------------- *)
-
-fun beginOpt f p (s : string, l : string list) : unit = f (p s) (s,l);
-
-fun endOpt () (_ : string, [] : string list) = ()
- | endOpt _ (_, _ :: _) = raise Bug "endOpt";
-
-fun stringOpt _ _ (_ : string, []) = raise Bug "stringOpt"
- | stringOpt f p (s, (h : string) :: t) : unit = f (p h) (s,t);
-
-local
- fun range NONE NONE = "Z"
- | range (SOME i) NONE = "{n IN Z | " ^ Int.toString i ^ " <= n}"
- | range NONE (SOME j) = "{n IN Z | n <= " ^ Int.toString j ^ "}"
- | range (SOME i) (SOME j) =
- "{n IN Z | " ^ Int.toString i ^ " <= n <= " ^ Int.toString j ^ "}";
- fun oLeq (SOME x) (SOME y) = x <= y | oLeq _ _ = true;
- fun argToInt arg omin omax x =
- (case Int.fromString x of
- SOME i =>
- if oLeq omin (SOME i) andalso oLeq (SOME i) omax then i else
- raise OptionExit
- {success = false, usage = false, message =
- SOME (arg ^ " option needs an integer argument in the range "
- ^ range omin omax ^ " (not " ^ x ^ ")")}
- | NONE =>
- raise OptionExit
- {success = false, usage = false, message =
- SOME (arg ^ " option needs an integer argument (not \"" ^ x ^ "\")")})
- handle Overflow =>
- raise OptionExit
- {success = false, usage = false, message =
- SOME (arg ^ " option suffered integer overflow on argument " ^ x)};
-in
- fun intOpt _ _ _ (_,[]) = raise Bug "intOpt"
- | intOpt (omin,omax) f p (s:string, h :: (t : string list)) : unit =
- f (p (argToInt s omin omax h)) (s,t);
-end;
-
-local
- fun range NONE NONE = "R"
- | range (SOME i) NONE = "{n IN R | " ^ Real.toString i ^ " <= n}"
- | range NONE (SOME j) = "{n IN R | n <= " ^ Real.toString j ^ "}"
- | range (SOME i) (SOME j) =
- "{n IN R | " ^ Real.toString i ^ " <= n <= " ^ Real.toString j ^ "}";
- fun oLeq (SOME (x:real)) (SOME y) = x <= y | oLeq _ _ = true;
- fun argToReal arg omin omax x =
- (case Real.fromString x of
- SOME i =>
- if oLeq omin (SOME i) andalso oLeq (SOME i) omax then i else
- raise OptionExit
- {success = false, usage = false, message =
- SOME (arg ^ " option needs an real argument in the range "
- ^ range omin omax ^ " (not " ^ x ^ ")")}
- | NONE =>
- raise OptionExit
- {success = false, usage = false, message =
- SOME (arg ^ " option needs an real argument (not \"" ^ x ^ "\")")})
-in
- fun realOpt _ _ _ (_,[]) = raise Bug "realOpt"
- | realOpt (omin,omax) f p (s:string, h :: (t : string list)) : unit =
- f (p (argToReal s omin omax h)) (s,t);
-end;
-
-fun enumOpt _ _ _ (_,[]) = raise Bug "enumOpt"
- | enumOpt (choices : string list) f p (s : string, h :: t) : unit =
- if mem h choices then f (p h) (s,t) else
- raise OptionExit
- {success = false, usage = false,
- message = SOME ("follow parameter " ^ s ^ " with one of {" ^
- join "," choices ^ "}, not \"" ^ h ^ "\"")};
-
-fun optionOpt _ _ _ (_,[]) = raise Bug "optionOpt"
- | optionOpt (x : string, p) f q (s : string, l as h :: t) : unit =
- if h = x then f (q NONE) (s,t) else p f (q o SOME) (s,l);
-
-(* ------------------------------------------------------------------------- *)
-(* Basic options useful for all programs *)
-(* ------------------------------------------------------------------------- *)
-
-val basicOptions : opt list =
- [{switches = ["--"], arguments = [],
- description = "no more options",
- processor = fn _ => raise Fail "basicOptions: --"},
- {switches = ["-?","-h","--help"], arguments = [],
- description = "display all options and exit",
- processor = fn _ => raise OptionExit
- {message = SOME "displaying all options", usage = true, success = true}},
- {switches = ["-v", "--version"], arguments = [],
- description = "display version information",
- processor = fn _ => raise Fail "basicOptions: -v, --version"}];
-
-(* ------------------------------------------------------------------------- *)
-(* All the command line options of a program *)
-(* ------------------------------------------------------------------------- *)
-
-type allOptions = {name : string, version : string, header : string,
- footer : string, options : opt list};
-
-(* ------------------------------------------------------------------------- *)
-(* Usage information *)
-(* ------------------------------------------------------------------------- *)
-
-fun versionInformation ({version, ...} : allOptions) = version;
-
-fun usageInformation ({name,version,header,footer,options} : allOptions) =
- let
- fun listOpts {switches = n, arguments = r, description = s,
- processor = _} =
- let
- fun indent (s, "" :: l) = indent (s ^ " ", l) | indent x = x
- val (res,n) = indent (" ",n)
- val res = res ^ join ", " n
- val res = foldl (fn (x,y) => y ^ " " ^ x) res r
- in
- [res ^ " ...", " " ^ s]
- end
-
- val alignment =
- [{leftAlign = true, padChar = #"."},
- {leftAlign = true, padChar = #" "}]
-
- val table = alignTable alignment (map listOpts options)
- in
- header ^ join "\n" table ^ "\n" ^ footer
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Exit the program gracefully *)
-(* ------------------------------------------------------------------------- *)
-
-fun exit (allopts : allOptions) (optexit : optionExit) =
- let
- val {name, options, ...} = allopts
- val {message, usage, success} = optexit
- fun err s = TextIO.output (TextIO.stdErr, s)
- in
- case message of NONE => () | SOME m => err (name ^ ": " ^ m ^ "\n");
- if usage then err (usageInformation allopts) else ();
- OS.Process.exit (if success then OS.Process.success else OS.Process.failure)
- end;
-
-fun succeed allopts =
- exit allopts {message = NONE, usage = false, success = true};
-
-fun fail allopts mesg =
- exit allopts {message = SOME mesg, usage = false, success = false};
-
-fun usage allopts mesg =
- exit allopts {message = SOME mesg, usage = true, success = false};
-
-fun version allopts =
- (print (versionInformation allopts);
- exit allopts {message = NONE, usage = false, success = true});
-
-(* ------------------------------------------------------------------------- *)
-(* Process the command line options passed to the program *)
-(* ------------------------------------------------------------------------- *)
-
-fun processOptions (allopts : allOptions) =
- let
- fun findOption x =
- case List.find (fn {switches = n, ...} => mem x n) (#options allopts) of
- NONE => raise OptionExit
- {message = SOME ("unknown switch \"" ^ x ^ "\""),
- usage = true, success = false}
- | SOME {arguments = r, processor = f, ...} => (r,f)
-
- fun getArgs x r xs =
- let
- fun f 1 = "a following argument"
- | f m = Int.toString m ^ " following arguments"
- val m = length r
- val () =
- if m <= length xs then () else
- raise OptionExit
- {usage = false, success = false, message = SOME
- (x ^ " option needs " ^ f m ^ ": " ^ join " " r)}
- in
- divide xs m
- end
-
- fun process [] = ([], [])
- | process ("--" :: xs) = ([("--",[])], xs)
- | process ("-v" :: _) = version allopts
- | process ("--version" :: _) = version allopts
- | process (x :: xs) =
- if x = "" orelse x = "-" orelse hd (explode x) <> #"-" then ([], x :: xs)
- else
- let
- val (r,f) = findOption x
- val (ys,xs) = getArgs x r xs
- val () = f (x,ys)
- in
- (cons (x,ys) ## I) (process xs)
- end
- in
- fn l =>
- let
- val (a,b) = process l
- val a = foldl (fn ((x,xs),ys) => x :: xs @ ys) [] (rev a)
- in
- (a,b)
- end
- handle OptionExit x => exit allopts x
- end;
-
-end
--- a/src/Tools/Metis/src/Ordered.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-(* ========================================================================= *)
-(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Ordered =
-sig
-
-type t
-
-val compare : t * t -> order
-
-(*
- PROVIDES
-
- !x : t. compare (x,x) = EQUAL
-
- !x y : t. compare (x,y) = LESS <=> compare (y,x) = GREATER
-
- !x y : t. compare (x,y) = EQUAL ==> compare (y,x) = EQUAL
-
- !x y z : t. compare (x,y) = EQUAL ==> compare (x,z) = compare (y,z)
-
- !x y z : t.
- compare (x,y) = LESS andalso compare (y,z) = LESS ==>
- compare (x,z) = LESS
-
- !x y z : t.
- compare (x,y) = GREATER andalso compare (y,z) = GREATER ==>
- compare (x,z) = GREATER
-*)
-
-end
--- a/src/Tools/Metis/src/Ordered.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(* ========================================================================= *)
-(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure IntOrdered =
-struct type t = int val compare = Int.compare end;
-
-structure StringOrdered =
-struct type t = string val compare = String.compare end;
--- a/src/Tools/Metis/src/PP.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,224 +0,0 @@
-(* ========================================================================= *)
-(* PP -- pretty-printing -- from the SML/NJ library *)
-(* *)
-(* Modified for Moscow ML from SML/NJ Library version 0.2 *)
-(* *)
-(* COPYRIGHT (c) 1992 by AT&T Bell Laboratories. *)
-(* *)
-(* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER. *)
-(* *)
-(* Permission to use, copy, modify, and distribute this software and its *)
-(* documentation for any purpose and without fee is hereby granted, *)
-(* provided that the above copyright notice appear in all copies and that *)
-(* both the copyright notice and this permission notice and warranty *)
-(* disclaimer appear in supporting documentation, and that the name of *)
-(* AT&T Bell Laboratories or any AT&T entity not be used in advertising *)
-(* or publicity pertaining to distribution of the software without *)
-(* specific, written prior permission. *)
-(* *)
-(* AT&T disclaims all warranties with regard to this software, including *)
-(* all implied warranties of merchantability and fitness. In no event *)
-(* shall AT&T be liable for any special, indirect or consequential *)
-(* damages or any damages whatsoever resulting from loss of use, data or *)
-(* profits, whether in an action of contract, negligence or other *)
-(* tortious action, arising out of or in connection with the use or *)
-(* performance of this software. *)
-(* ========================================================================= *)
-
-signature PP =
-sig
-
- type ppstream
-
- type ppconsumer =
- {consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit}
-
- datatype break_style =
- CONSISTENT
- | INCONSISTENT
-
- val mk_ppstream : ppconsumer -> ppstream
-
- val dest_ppstream : ppstream -> ppconsumer
-
- val add_break : ppstream -> int * int -> unit
-
- val add_newline : ppstream -> unit
-
- val add_string : ppstream -> string -> unit
-
- val begin_block : ppstream -> break_style -> int -> unit
-
- val end_block : ppstream -> unit
-
- val clear_ppstream : ppstream -> unit
-
- val flush_ppstream : ppstream -> unit
-
- val with_pp : ppconsumer -> (ppstream -> unit) -> unit
-
- val pp_to_string : int -> (ppstream -> 'a -> unit) -> 'a -> string
-
-end
-
-(*
- This structure provides tools for creating customized Oppen-style
- pretty-printers, based on the type ppstream. A ppstream is an
- output stream that contains prettyprinting commands. The commands
- are placed in the stream by various function calls listed below.
-
- There following primitives add commands to the stream:
- begin_block, end_block, add_string, add_break, and add_newline.
- All calls to add_string, add_break, and add_newline must happen
- between a pair of calls to begin_block and end_block must be
- properly nested dynamically. All calls to begin_block and
- end_block must be properly nested (dynamically).
-
- [ppconsumer] is the type of sinks for pretty-printing. A value of
- type ppconsumer is a record
- { consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit }
- of a string consumer, a specified linewidth, and a flush function
- which is called whenever flush_ppstream is called.
-
- A prettyprinter can be called outright to print a value. In
- addition, a prettyprinter for a base type or nullary datatype ty
- can be installed in the top-level system. Then the installed
- prettyprinter will be invoked automatically whenever a value of
- type ty is to be printed.
-
- [break_style] is the type of line break styles for blocks:
-
- [CONSISTENT] specifies that if any line break occurs inside the
- block, then all indicated line breaks occur.
-
- [INCONSISTENT] specifies that breaks will be inserted to only to
- avoid overfull lines.
-
- [mk_ppstream {consumer, linewidth, flush}] creates a new ppstream
- which invokes the consumer to output text, putting at most
- linewidth characters on each line.
-
- [dest_ppstream ppstrm] extracts the linewidth, flush function, and
- consumer from a ppstream.
-
- [add_break ppstrm (size, offset)] notifies the pretty-printer that
- a line break is possible at this point.
- * When the current block style is CONSISTENT:
- ** if the entire block fits on the remainder of the line, then
- output size spaces; else
- ** increase the current indentation by the block offset;
- further indent every item of the block by offset, and add
- one newline at every add_break in the block.
- * When the current block style is INCONSISTENT:
- ** if the next component of the block fits on the remainder of
- the line, then output size spaces; else
- ** issue a newline and indent to the current indentation level
- plus the block offset plus the offset.
-
- [add_newline ppstrm] issues a newline.
-
- [add_string ppstrm str] outputs the string str to the ppstream.
-
- [begin_block ppstrm style blockoffset] begins a new block and
- level of indentation, with the given style and block offset.
-
- [end_block ppstrm] closes the current block.
-
- [clear_ppstream ppstrm] restarts the stream, without affecting the
- underlying consumer.
-
- [flush_ppstream ppstrm] executes any remaining commands in the
- ppstream (that is, flushes currently accumulated output to the
- consumer associated with ppstrm); executes the flush function
- associated with the consumer; and calls clear_ppstream.
-
- [with_pp consumer f] makes a new ppstream from the consumer and
- applies f (which can be thought of as a producer) to that
- ppstream, then flushed the ppstream and returns the value of f.
-
- [pp_to_string linewidth printit x] constructs a new ppstream
- ppstrm whose consumer accumulates the output in a string s. Then
- evaluates (printit ppstrm x) and finally returns the string s.
-
-
- Example 1: A simple prettyprinter for Booleans:
-
- load "PP";
- fun ppbool pps d =
- let open PP
- in
- begin_block pps INCONSISTENT 6;
- add_string pps (if d then "right" else "wrong");
- end_block pps
- end;
-
- Now one may define a ppstream to print to, and exercise it:
-
- val ppstrm = PP.mk_ppstream {consumer =
- fn s => TextIO.output(TextIO.stdOut, s),
- linewidth = 72,
- flush =
- fn () => TextIO.flushOut TextIO.stdOut};
-
- fun ppb b = (ppbool ppstrm b; PP.flush_ppstream ppstrm);
-
- - ppb false;
- wrong> val it = () : unit
-
- The prettyprinter may also be installed in the toplevel system;
- then it will be used to print all expressions of type bool
- subsequently computed:
-
- - installPP ppbool;
- > val it = () : unit
- - 1=0;
- > val it = wrong : bool
- - 1=1;
- > val it = right : bool
-
- See library Meta for a description of installPP.
-
-
- Example 2: Prettyprinting simple expressions (examples/pretty/ppexpr.sml):
-
- datatype expr =
- Cst of int
- | Neg of expr
- | Plus of expr * expr
-
- fun ppexpr pps e0 =
- let open PP
- fun ppe (Cst i) = add_string pps (Int.toString i)
- | ppe (Neg e) = (add_string pps "~"; ppe e)
- | ppe (Plus(e1, e2)) = (begin_block pps CONSISTENT 0;
- add_string pps "(";
- ppe e1;
- add_string pps " + ";
- add_break pps (0, 1);
- ppe e2;
- add_string pps ")";
- end_block pps)
- in
- begin_block pps INCONSISTENT 0;
- ppe e0;
- end_block pps
- end
-
- val _ = installPP ppexpr;
-
- (* Some example values: *)
-
- val e1 = Cst 1;
- val e2 = Cst 2;
- val e3 = Plus(e1, Neg e2);
- val e4 = Plus(Neg e3, e3);
- val e5 = Plus(Neg e4, e4);
- val e6 = Plus(e5, e5);
- val e7 = Plus(e6, e6);
- val e8 =
- Plus(e3, Plus(e3, Plus(e3, Plus(e3, Plus(e3, Plus(e3, e7))))));
-*)
--- a/src/Tools/Metis/src/PP.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,617 +0,0 @@
-(* ========================================================================= *)
-(* PP -- pretty-printing -- from the SML/NJ library *)
-(* *)
-(* Modified for Moscow ML from SML/NJ Library version 0.2 *)
-(* *)
-(* COPYRIGHT (c) 1992 by AT&T Bell Laboratories. *)
-(* *)
-(* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER. *)
-(* *)
-(* Permission to use, copy, modify, and distribute this software and its *)
-(* documentation for any purpose and without fee is hereby granted, *)
-(* provided that the above copyright notice appear in all copies and that *)
-(* both the copyright notice and this permission notice and warranty *)
-(* disclaimer appear in supporting documentation, and that the name of *)
-(* AT&T Bell Laboratories or any AT&T entity not be used in advertising *)
-(* or publicity pertaining to distribution of the software without *)
-(* specific, written prior permission. *)
-(* *)
-(* AT&T disclaims all warranties with regard to this software, including *)
-(* all implied warranties of merchantability and fitness. In no event *)
-(* shall AT&T be liable for any special, indirect or consequential *)
-(* damages or any damages whatsoever resulting from loss of use, data or *)
-(* profits, whether in an action of contract, negligence or other *)
-(* tortious action, arising out of or in connection with the use or *)
-(* performance of this software. *)
-(* ========================================================================= *)
-
-structure PP :> PP =
-struct
-
-open Array
-infix 9 sub
-
-(* the queue library, formerly in unit Ppqueue *)
-
-datatype Qend = Qback | Qfront
-
-exception QUEUE_FULL
-exception QUEUE_EMPTY
-exception REQUESTED_QUEUE_SIZE_TOO_SMALL
-
-local
- fun ++ i n = (i + 1) mod n
- fun -- i n = (i - 1) mod n
-in
-
-abstype 'a queue = QUEUE of {elems: 'a array, (* the contents *)
- front: int ref,
- back: int ref,
- size: int} (* fixed size of element array *)
-with
-
- fun is_empty (QUEUE{front=ref ~1, back=ref ~1,...}) = true
- | is_empty _ = false
-
- fun mk_queue n init_val =
- if (n < 2)
- then raise REQUESTED_QUEUE_SIZE_TOO_SMALL
- else QUEUE{elems=array(n, init_val), front=ref ~1, back=ref ~1, size=n}
-
- fun clear_queue (QUEUE{front,back,...}) = (front := ~1; back := ~1)
-
- fun queue_at Qfront (QUEUE{elems,front,...}) = elems sub !front
- | queue_at Qback (QUEUE{elems,back,...}) = elems sub !back
-
- fun en_queue Qfront item (Q as QUEUE{elems,front,back,size}) =
- if (is_empty Q)
- then (front := 0; back := 0;
- update(elems,0,item))
- else let val i = --(!front) size
- in if (i = !back)
- then raise QUEUE_FULL
- else (update(elems,i,item); front := i)
- end
- | en_queue Qback item (Q as QUEUE{elems,front,back,size}) =
- if (is_empty Q)
- then (front := 0; back := 0;
- update(elems,0,item))
- else let val i = ++(!back) size
- in if (i = !front)
- then raise QUEUE_FULL
- else (update(elems,i,item); back := i)
- end
-
- fun de_queue Qfront (Q as QUEUE{front,back,size,...}) =
- if (!front = !back) (* unitary queue *)
- then clear_queue Q
- else front := ++(!front) size
- | de_queue Qback (Q as QUEUE{front,back,size,...}) =
- if (!front = !back)
- then clear_queue Q
- else back := --(!back) size
-
-end (* abstype queue *)
-end (* local *)
-
-
-val magic: 'a -> 'a = fn x => x
-
-(* exception PP_FAIL of string *)
-
-datatype break_style = CONSISTENT | INCONSISTENT
-
-datatype break_info
- = FITS
- | PACK_ONTO_LINE of int
- | ONE_PER_LINE of int
-
-(* Some global values *)
-val INFINITY = 999999
-
-abstype indent_stack = Istack of break_info list ref
-with
- fun mk_indent_stack() = Istack (ref([]:break_info list))
- fun clear_indent_stack (Istack stk) = (stk := ([]:break_info list))
- fun top (Istack stk) =
- case !stk
- of nil => raise Fail "PP-error: top: badly formed block"
- | x::_ => x
- fun push (x,(Istack stk)) = stk := x::(!stk)
- fun pop (Istack stk) =
- case !stk
- of nil => raise Fail "PP-error: pop: badly formed block"
- | _::rest => stk := rest
-end
-
-(* The delim_stack is used to compute the size of blocks. It is
- a stack of indices into the token buffer. The indices only point to
- BBs, Es, and BRs. We push BBs and Es onto the stack until a BR
- is encountered. Then we compute sizes and pop. When we encounter
- a BR in the middle of a block, we compute the Distance_to_next_break
- of the previous BR in the block, if there was one.
-
- We need to be able to delete from the bottom of the delim_stack, so
- we use a queue, treated with a stack discipline, i.e., we only add
- items at the head of the queue, but can delete from the front or
- back of the queue.
-*)
-abstype delim_stack = Dstack of int queue
-with
- fun new_delim_stack i = Dstack(mk_queue i ~1)
- fun reset_delim_stack (Dstack q) = clear_queue q
-
- fun pop_delim_stack (Dstack d) = de_queue Qfront d
- fun pop_bottom_delim_stack (Dstack d) = de_queue Qback d
-
- fun push_delim_stack(i,Dstack d) = en_queue Qfront i d
- fun top_delim_stack (Dstack d) = queue_at Qfront d
- fun bottom_delim_stack (Dstack d) = queue_at Qback d
- fun delim_stack_is_empty (Dstack d) = is_empty d
-end
-
-
-type block_info = { Block_size : int ref,
- Block_offset : int,
- How_to_indent : break_style }
-
-
-(* Distance_to_next_break includes Number_of_blanks. Break_offset is
- a local offset for the break. BB represents a sequence of contiguous
- Begins. E represents a sequence of contiguous Ends.
-*)
-datatype pp_token
- = S of {String : string, Length : int}
- | BB of {Pblocks : block_info list ref, (* Processed *)
- Ublocks : block_info list ref} (* Unprocessed *)
- | E of {Pend : int ref, Uend : int ref}
- | BR of {Distance_to_next_break : int ref,
- Number_of_blanks : int,
- Break_offset : int}
-
-
-(* The initial values in the token buffer *)
-val initial_token_value = S{String = "", Length = 0}
-
-(* type ppstream = General.ppstream; *)
-datatype ppstream_ =
- PPS of
- {consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit,
- the_token_buffer : pp_token array,
- the_delim_stack : delim_stack,
- the_indent_stack : indent_stack,
- ++ : int ref -> unit, (* increment circular buffer index *)
- space_left : int ref, (* remaining columns on page *)
- left_index : int ref, (* insertion index *)
- right_index : int ref, (* output index *)
- left_sum : int ref, (* size of strings and spaces inserted *)
- right_sum : int ref} (* size of strings and spaces printed *)
-
-type ppstream = ppstream_
-
-type ppconsumer = {consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit}
-
-fun mk_ppstream {consumer,linewidth,flush} =
- if (linewidth<5)
- then raise Fail "PP-error: linewidth too_small"
- else let val buf_size = 3*linewidth
- in magic(
- PPS{consumer = consumer,
- linewidth = linewidth,
- flush = flush,
- the_token_buffer = array(buf_size, initial_token_value),
- the_delim_stack = new_delim_stack buf_size,
- the_indent_stack = mk_indent_stack (),
- ++ = fn i => i := ((!i + 1) mod buf_size),
- space_left = ref linewidth,
- left_index = ref 0, right_index = ref 0,
- left_sum = ref 0, right_sum = ref 0}
- ) : ppstream
- end
-
-fun dest_ppstream(pps : ppstream) =
- let val PPS{consumer,linewidth,flush, ...} = magic pps
- in {consumer=consumer,linewidth=linewidth,flush=flush} end
-
-local
- val space = " "
- fun mk_space (0,s) = String.concat s
- | mk_space (n,s) = mk_space((n-1), (space::s))
- val space_table = Vector.tabulate(100, fn i => mk_space(i,[]))
- fun nspaces n = Vector.sub(space_table, n)
- handle General.Subscript =>
- if n < 0
- then ""
- else let val n2 = n div 2
- val n2_spaces = nspaces n2
- val extra = if (n = (2*n2)) then "" else space
- in String.concat [n2_spaces, n2_spaces, extra]
- end
-in
- fun cr_indent (ofn, i) = ofn ("\n"^(nspaces i))
- fun indent (ofn,i) = ofn (nspaces i)
-end
-
-
-(* Print a the first member of a contiguous sequence of Begins. If there
- are "processed" Begins, then take the first off the list. If there are
- no processed Begins, take the last member off the "unprocessed" list.
- This works because the unprocessed list is treated as a stack, the
- processed list as a FIFO queue. How can an item on the unprocessed list
- be printable? Because of what goes on in add_string. See there for details.
-*)
-
-fun print_BB (_,{Pblocks = ref [], Ublocks = ref []}) =
- raise Fail "PP-error: print_BB"
- | print_BB (PPS{the_indent_stack,linewidth,space_left=ref sp_left,...},
- {Pblocks as ref({How_to_indent=CONSISTENT,Block_size,
- Block_offset}::rst),
- Ublocks=ref[]}) =
- (push ((if (!Block_size > sp_left)
- then ONE_PER_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- Pblocks := rst)
- | print_BB(PPS{the_indent_stack,linewidth,space_left=ref sp_left,...},
- {Pblocks as ref({Block_size,Block_offset,...}::rst),Ublocks=ref[]}) =
- (push ((if (!Block_size > sp_left)
- then PACK_ONTO_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- Pblocks := rst)
- | print_BB (PPS{the_indent_stack, linewidth, space_left=ref sp_left,...},
- {Ublocks,...}) =
- let fun pr_end_Ublock [{How_to_indent=CONSISTENT,Block_size,Block_offset}] l =
- (push ((if (!Block_size > sp_left)
- then ONE_PER_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- List.rev l)
- | pr_end_Ublock [{Block_size,Block_offset,...}] l =
- (push ((if (!Block_size > sp_left)
- then PACK_ONTO_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- List.rev l)
- | pr_end_Ublock (a::rst) l = pr_end_Ublock rst (a::l)
- | pr_end_Ublock _ _ =
- raise Fail "PP-error: print_BB: internal error"
- in Ublocks := pr_end_Ublock(!Ublocks) []
- end
-
-
-(* Uend should always be 0 when print_E is called. *)
-fun print_E (_,{Pend = ref 0, Uend = ref 0}) =
- raise Fail "PP-error: print_E"
- | print_E (istack,{Pend, ...}) =
- let fun pop_n_times 0 = ()
- | pop_n_times n = (pop istack; pop_n_times(n-1))
- in pop_n_times(!Pend); Pend := 0
- end
-
-
-(* "cursor" is how many spaces across the page we are. *)
-
-fun print_token(PPS{consumer,space_left,...}, S{String,Length}) =
- (consumer String;
- space_left := (!space_left) - Length)
- | print_token(ppstrm,BB b) = print_BB(ppstrm,b)
- | print_token(PPS{the_indent_stack,...},E e) =
- print_E (the_indent_stack,e)
- | print_token (PPS{the_indent_stack,space_left,consumer,linewidth,...},
- BR{Distance_to_next_break,Number_of_blanks,Break_offset}) =
- (case (top the_indent_stack)
- of FITS =>
- (space_left := (!space_left) - Number_of_blanks;
- indent (consumer,Number_of_blanks))
- | (ONE_PER_LINE cursor) =>
- let val new_cursor = cursor + Break_offset
- in space_left := linewidth - new_cursor;
- cr_indent (consumer,new_cursor)
- end
- | (PACK_ONTO_LINE cursor) =>
- if (!Distance_to_next_break > (!space_left))
- then let val new_cursor = cursor + Break_offset
- in space_left := linewidth - new_cursor;
- cr_indent(consumer,new_cursor)
- end
- else (space_left := !space_left - Number_of_blanks;
- indent (consumer,Number_of_blanks)))
-
-
-fun clear_ppstream(pps : ppstream) =
- let val PPS{the_token_buffer, the_delim_stack,
- the_indent_stack,left_sum, right_sum,
- left_index, right_index,space_left,linewidth,...}
- = magic pps
- val buf_size = 3*linewidth
- fun set i =
- if (i = buf_size)
- then ()
- else (update(the_token_buffer,i,initial_token_value);
- set (i+1))
- in set 0;
- clear_indent_stack the_indent_stack;
- reset_delim_stack the_delim_stack;
- left_sum := 0; right_sum := 0;
- left_index := 0; right_index := 0;
- space_left := linewidth
- end
-
-
-(* Move insertion head to right unless adding a BB and already at a BB,
- or unless adding an E and already at an E.
-*)
-fun BB_inc_right_index(PPS{the_token_buffer, right_index, ++,...})=
- case (the_token_buffer sub (!right_index))
- of (BB _) => ()
- | _ => ++right_index
-
-fun E_inc_right_index(PPS{the_token_buffer,right_index, ++,...})=
- case (the_token_buffer sub (!right_index))
- of (E _) => ()
- | _ => ++right_index
-
-
-fun pointers_coincide(PPS{left_index,right_index,the_token_buffer,...}) =
- (!left_index = !right_index) andalso
- (case (the_token_buffer sub (!left_index))
- of (BB {Pblocks = ref [], Ublocks = ref []}) => true
- | (BB _) => false
- | (E {Pend = ref 0, Uend = ref 0}) => true
- | (E _) => false
- | _ => true)
-
-fun advance_left (ppstrm as PPS{consumer,left_index,left_sum,
- the_token_buffer,++,...},
- instr) =
- let val NEG = ~1
- val POS = 0
- fun inc_left_sum (BR{Number_of_blanks, ...}) =
- left_sum := (!left_sum) + Number_of_blanks
- | inc_left_sum (S{Length, ...}) = left_sum := (!left_sum) + Length
- | inc_left_sum _ = ()
-
- fun last_size [{Block_size, ...}:block_info] = !Block_size
- | last_size (_::rst) = last_size rst
- | last_size _ = raise Fail "PP-error: last_size: internal error"
- fun token_size (S{Length, ...}) = Length
- | token_size (BB b) =
- (case b
- of {Pblocks = ref [], Ublocks = ref []} =>
- raise Fail "PP-error: BB_size"
- | {Pblocks as ref(_::_),Ublocks=ref[]} => POS
- | {Ublocks, ...} => last_size (!Ublocks))
- | token_size (E{Pend = ref 0, Uend = ref 0}) =
- raise Fail "PP-error: token_size.E"
- | token_size (E{Pend = ref 0, ...}) = NEG
- | token_size (E _) = POS
- | token_size (BR {Distance_to_next_break, ...}) = !Distance_to_next_break
- fun loop (instr) =
- if (token_size instr < 0) (* synchronization point; cannot advance *)
- then ()
- else (print_token(ppstrm,instr);
- inc_left_sum instr;
- if (pointers_coincide ppstrm)
- then ()
- else (* increment left index *)
-
- (* When this is evaluated, we know that the left_index has not yet
- caught up to the right_index. If we are at a BB or an E, we can
- increment left_index if there is no work to be done, i.e., all Begins
- or Ends have been dealt with. Also, we should do some housekeeping and
- clear the buffer at left_index, otherwise we can get errors when
- left_index catches up to right_index and we reset the indices to 0.
- (We might find ourselves adding a BB to an "old" BB, with the result
- that the index is not pushed onto the delim_stack. This can lead to
- mangled output.)
- *)
- (case (the_token_buffer sub (!left_index))
- of (BB {Pblocks = ref [], Ublocks = ref []}) =>
- (update(the_token_buffer,!left_index,
- initial_token_value);
- ++left_index)
- | (BB _) => ()
- | (E {Pend = ref 0, Uend = ref 0}) =>
- (update(the_token_buffer,!left_index,
- initial_token_value);
- ++left_index)
- | (E _) => ()
- | _ => ++left_index;
- loop (the_token_buffer sub (!left_index))))
- in loop instr
- end
-
-
-fun begin_block (pps : ppstream) style offset =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer, the_delim_stack,left_index,
- left_sum, right_index, right_sum,...}
- = ppstrm
- in
- (if (delim_stack_is_empty the_delim_stack)
- then (left_index := 0;
- left_sum := 1;
- right_index := 0;
- right_sum := 1)
- else BB_inc_right_index ppstrm;
- case (the_token_buffer sub (!right_index))
- of (BB {Ublocks, ...}) =>
- Ublocks := {Block_size = ref (~(!right_sum)),
- Block_offset = offset,
- How_to_indent = style}::(!Ublocks)
- | _ => (update(the_token_buffer, !right_index,
- BB{Pblocks = ref [],
- Ublocks = ref [{Block_size = ref (~(!right_sum)),
- Block_offset = offset,
- How_to_indent = style}]});
- push_delim_stack (!right_index, the_delim_stack)))
- end
-
-fun end_block(pps : ppstream) =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer,the_delim_stack,right_index,...}
- = ppstrm
- in
- if (delim_stack_is_empty the_delim_stack)
- then print_token(ppstrm,(E{Pend = ref 1, Uend = ref 0}))
- else (E_inc_right_index ppstrm;
- case (the_token_buffer sub (!right_index))
- of (E{Uend, ...}) => Uend := !Uend + 1
- | _ => (update(the_token_buffer,!right_index,
- E{Uend = ref 1, Pend = ref 0});
- push_delim_stack (!right_index, the_delim_stack)))
- end
-
-local
- fun check_delim_stack(PPS{the_token_buffer,the_delim_stack,right_sum,...}) =
- let fun check k =
- if (delim_stack_is_empty the_delim_stack)
- then ()
- else case(the_token_buffer sub (top_delim_stack the_delim_stack))
- of (BB{Ublocks as ref ((b as {Block_size, ...})::rst),
- Pblocks}) =>
- if (k>0)
- then (Block_size := !right_sum + !Block_size;
- Pblocks := b :: (!Pblocks);
- Ublocks := rst;
- if (List.length rst = 0)
- then pop_delim_stack the_delim_stack
- else ();
- check(k-1))
- else ()
- | (E{Pend,Uend}) =>
- (Pend := (!Pend) + (!Uend);
- Uend := 0;
- pop_delim_stack the_delim_stack;
- check(k + !Pend))
- | (BR{Distance_to_next_break, ...}) =>
- (Distance_to_next_break :=
- !right_sum + !Distance_to_next_break;
- pop_delim_stack the_delim_stack;
- if (k>0)
- then check k
- else ())
- | _ => raise Fail "PP-error: check_delim_stack.catchall"
- in check 0
- end
-in
-
- fun add_break (pps : ppstream) (n, break_offset) =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer,the_delim_stack,left_index,
- right_index,left_sum,right_sum, ++, ...}
- = ppstrm
- in
- (if (delim_stack_is_empty the_delim_stack)
- then (left_index := 0; right_index := 0;
- left_sum := 1; right_sum := 1)
- else ++right_index;
- update(the_token_buffer, !right_index,
- BR{Distance_to_next_break = ref (~(!right_sum)),
- Number_of_blanks = n,
- Break_offset = break_offset});
- check_delim_stack ppstrm;
- right_sum := (!right_sum) + n;
- push_delim_stack (!right_index,the_delim_stack))
- end
-
- fun flush_ppstream0(pps : ppstream) =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_delim_stack,the_token_buffer, flush, left_index,...}
- = ppstrm
- in
- (if (delim_stack_is_empty the_delim_stack)
- then ()
- else (check_delim_stack ppstrm;
- advance_left(ppstrm, the_token_buffer sub (!left_index)));
- flush())
- end
-
-end (* local *)
-
-
-fun flush_ppstream ppstrm =
- (flush_ppstream0 ppstrm;
- clear_ppstream ppstrm)
-
-fun add_string (pps : ppstream) s =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer,the_delim_stack,consumer,
- right_index,right_sum,left_sum,
- left_index,space_left,++,...}
- = ppstrm
- fun fnl [{Block_size, ...}:block_info] = Block_size := INFINITY
- | fnl (_::rst) = fnl rst
- | fnl _ = raise Fail "PP-error: fnl: internal error"
-
- fun set(dstack,BB{Ublocks as ref[{Block_size,...}:block_info],...}) =
- (pop_bottom_delim_stack dstack;
- Block_size := INFINITY)
- | set (_,BB {Ublocks = ref(_::rst), ...}) = fnl rst
- | set (dstack, E{Pend,Uend}) =
- (Pend := (!Pend) + (!Uend);
- Uend := 0;
- pop_bottom_delim_stack dstack)
- | set (dstack,BR{Distance_to_next_break,...}) =
- (pop_bottom_delim_stack dstack;
- Distance_to_next_break := INFINITY)
- | set _ = raise (Fail "PP-error: add_string.set")
-
- fun check_stream () =
- if ((!right_sum - !left_sum) > !space_left)
- then if (delim_stack_is_empty the_delim_stack)
- then ()
- else let val i = bottom_delim_stack the_delim_stack
- in if (!left_index = i)
- then set (the_delim_stack, the_token_buffer sub i)
- else ();
- advance_left(ppstrm,
- the_token_buffer sub (!left_index));
- if (pointers_coincide ppstrm)
- then ()
- else check_stream ()
- end
- else ()
-
- val slen = String.size s
- val S_token = S{String = s, Length = slen}
-
- in if (delim_stack_is_empty the_delim_stack)
- then print_token(ppstrm,S_token)
- else (++right_index;
- update(the_token_buffer, !right_index, S_token);
- right_sum := (!right_sum)+slen;
- check_stream ())
- end
-
-
-(* Derived form. The +2 is for peace of mind *)
-fun add_newline (pps : ppstream) =
- let val PPS{linewidth, ...} = magic pps
- in add_break pps (linewidth+2,0) end
-
-(* Derived form. Builds a ppstream, sends pretty printing commands called in
- f to the ppstream, then flushes ppstream.
-*)
-
-fun with_pp ppconsumer ppfn =
- let val ppstrm = mk_ppstream ppconsumer
- in ppfn ppstrm;
- flush_ppstream0 ppstrm
- end
- handle Fail msg =>
- (TextIO.print (">>>> Pretty-printer failure: " ^ msg ^ "\n"))
-
-fun pp_to_string linewidth ppfn ob =
- let val l = ref ([]:string list)
- fun attach s = l := (s::(!l))
- in with_pp {consumer = attach, linewidth=linewidth, flush = fn()=>()}
- (fn ppstrm => ppfn ppstrm ob);
- String.concat(List.rev(!l))
- end
-end
--- a/src/Tools/Metis/src/Parser.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,144 +0,0 @@
-(* ========================================================================= *)
-(* PARSING AND PRETTY PRINTING *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Parser =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty printing for built-in types *)
-(* ------------------------------------------------------------------------- *)
-
-type ppstream = PP.ppstream
-
-datatype breakStyle = Consistent | Inconsistent
-
-type 'a pp = ppstream -> 'a -> unit
-
-val lineLength : int ref
-
-val beginBlock : ppstream -> breakStyle -> int -> unit
-
-val endBlock : ppstream -> unit
-
-val addString : ppstream -> string -> unit
-
-val addBreak : ppstream -> int * int -> unit
-
-val addNewline : ppstream -> unit
-
-val ppMap : ('a -> 'b) -> 'b pp -> 'a pp
-
-val ppBracket : string -> string -> 'a pp -> 'a pp
-
-val ppSequence : string -> 'a pp -> 'a list pp
-
-val ppBinop : string -> 'a pp -> 'b pp -> ('a * 'b) pp
-
-val ppChar : char pp
-
-val ppString : string pp
-
-val ppUnit : unit pp
-
-val ppBool : bool pp
-
-val ppInt : int pp
-
-val ppReal : real pp
-
-val ppOrder : order pp
-
-val ppList : 'a pp -> 'a list pp
-
-val ppOption : 'a pp -> 'a option pp
-
-val ppPair : 'a pp -> 'b pp -> ('a * 'b) pp
-
-val ppTriple : 'a pp -> 'b pp -> 'c pp -> ('a * 'b * 'c) pp
-
-val toString : 'a pp -> 'a -> string (* Uses !lineLength *)
-
-val fromString : ('a -> string) -> 'a pp
-
-val ppTrace : 'a pp -> string -> 'a -> unit
-
-(* ------------------------------------------------------------------------- *)
-(* Recursive descent parsing combinators *)
-(* ------------------------------------------------------------------------- *)
-
-(* Generic parsers
-
-Recommended fixities:
- infixr 9 >>++
- infixr 8 ++
- infixr 7 >>
- infixr 6 ||
-*)
-
-exception NoParse
-
-val error : 'a -> 'b * 'a
-
-val ++ : ('a -> 'b * 'a) * ('a -> 'c * 'a) -> 'a -> ('b * 'c) * 'a
-
-val >> : ('a -> 'b * 'a) * ('b -> 'c) -> 'a -> 'c * 'a
-
-val >>++ : ('a -> 'b * 'a) * ('b -> 'a -> 'c * 'a) -> 'a -> 'c * 'a
-
-val || : ('a -> 'b * 'a) * ('a -> 'b * 'a) -> 'a -> 'b * 'a
-
-val first : ('a -> 'b * 'a) list -> 'a -> 'b * 'a
-
-val mmany : ('s -> 'a -> 's * 'a) -> 's -> 'a -> 's * 'a
-
-val many : ('a -> 'b * 'a) -> 'a -> 'b list * 'a
-
-val atLeastOne : ('a -> 'b * 'a) -> 'a -> 'b list * 'a
-
-val nothing : 'a -> unit * 'a
-
-val optional : ('a -> 'b * 'a) -> 'a -> 'b option * 'a
-
-(* Stream based parsers *)
-
-type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream
-
-val everything : ('a, 'b list) parser -> 'a Stream.stream -> 'b Stream.stream
-
-val maybe : ('a -> 'b option) -> ('a,'b) parser
-
-val finished : ('a,unit) parser
-
-val some : ('a -> bool) -> ('a,'a) parser
-
-val any : ('a,'a) parser
-
-val exact : ''a -> (''a,''a) parser
-
-(* ------------------------------------------------------------------------- *)
-(* Infix operators *)
-(* ------------------------------------------------------------------------- *)
-
-type infixities = {token : string, precedence : int, leftAssoc : bool} list
-
-val infixTokens : infixities -> string list
-
-val parseInfixes :
- infixities -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
- (string,'a) parser
-
-val ppInfixes :
- infixities -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp ->
- ('a * bool) pp
-
-(* ------------------------------------------------------------------------- *)
-(* Quotations *)
-(* ------------------------------------------------------------------------- *)
-
-type 'a quotation = 'a frag list
-
-val parseQuotation : ('a -> string) -> (string -> 'b) -> 'a quotation -> 'b
-
-end
--- a/src/Tools/Metis/src/Parser.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,397 +0,0 @@
-(* ========================================================================= *)
-(* PARSER COMBINATORS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Parser :> Parser =
-struct
-
-infixr 9 >>++
-infixr 8 ++
-infixr 7 >>
-infixr 6 ||
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions. *)
-(* ------------------------------------------------------------------------- *)
-
-exception Bug = Useful.Bug;
-
-val trace = Useful.trace
-and equal = Useful.equal
-and I = Useful.I
-and K = Useful.K
-and C = Useful.C
-and fst = Useful.fst
-and snd = Useful.snd
-and pair = Useful.pair
-and curry = Useful.curry
-and funpow = Useful.funpow
-and mem = Useful.mem
-and sortMap = Useful.sortMap;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty printing for built-in types *)
-(* ------------------------------------------------------------------------- *)
-
-type ppstream = PP.ppstream
-
-datatype breakStyle = Consistent | Inconsistent
-
-type 'a pp = PP.ppstream -> 'a -> unit;
-
-val lineLength = ref 75;
-
-fun beginBlock pp Consistent = PP.begin_block pp PP.CONSISTENT
- | beginBlock pp Inconsistent = PP.begin_block pp PP.INCONSISTENT;
-
-val endBlock = PP.end_block
-and addString = PP.add_string
-and addBreak = PP.add_break
-and addNewline = PP.add_newline;
-
-fun ppMap f ppA (ppstrm : PP.ppstream) x : unit = ppA ppstrm (f x);
-
-fun ppBracket l r ppA pp a =
- let
- val ln = size l
- in
- beginBlock pp Inconsistent ln;
- if ln = 0 then () else addString pp l;
- ppA pp a;
- if r = "" then () else addString pp r;
- endBlock pp
- end;
-
-fun ppSequence sep ppA pp =
- let
- fun ppX x = (addString pp sep; addBreak pp (1,0); ppA pp x)
- in
- fn [] => ()
- | h :: t =>
- (beginBlock pp Inconsistent 0;
- ppA pp h;
- app ppX t;
- endBlock pp)
- end;
-
-fun ppBinop s ppA ppB pp (a,b) =
- (beginBlock pp Inconsistent 0;
- ppA pp a;
- if s = "" then () else addString pp s;
- addBreak pp (1,0);
- ppB pp b;
- endBlock pp);
-
-fun ppTrinop ab bc ppA ppB ppC pp (a,b,c) =
- (beginBlock pp Inconsistent 0;
- ppA pp a;
- if ab = "" then () else addString pp ab;
- addBreak pp (1,0);
- ppB pp b;
- if bc = "" then () else addString pp bc;
- addBreak pp (1,0);
- ppC pp c;
- endBlock pp);
-
-(* Pretty-printers for common types *)
-
-fun ppString pp s =
- (beginBlock pp Inconsistent 0;
- addString pp s;
- endBlock pp);
-
-val ppUnit = ppMap (fn () => "()") ppString;
-
-val ppChar = ppMap str ppString;
-
-val ppBool = ppMap (fn true => "true" | false => "false") ppString;
-
-val ppInt = ppMap Int.toString ppString;
-
-val ppReal = ppMap Real.toString ppString;
-
-val ppOrder =
- let
- fun f LESS = "Less"
- | f EQUAL = "Equal"
- | f GREATER = "Greater"
- in
- ppMap f ppString
- end;
-
-fun ppList ppA = ppBracket "[" "]" (ppSequence "," ppA);
-
-fun ppOption _ pp NONE = ppString pp "-"
- | ppOption ppA pp (SOME a) = ppA pp a;
-
-fun ppPair ppA ppB = ppBracket "(" ")" (ppBinop "," ppA ppB);
-
-fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppTrinop "," "," ppA ppB ppC);
-
-(* Keep eta expanded to evaluate lineLength when supplied with a *)
-fun toString ppA a = PP.pp_to_string (!lineLength) ppA a;
-
-fun fromString toS = ppMap toS ppString;
-
-fun ppTrace ppX nameX x =
- trace (toString (ppBinop " =" ppString ppX) (nameX,x) ^ "\n");
-
-(* ------------------------------------------------------------------------- *)
-(* Generic. *)
-(* ------------------------------------------------------------------------- *)
-
-exception NoParse;
-
-val error : 'a -> 'b * 'a = fn _ => raise NoParse;
-
-fun op ++ (parser1,parser2) input =
- let
- val (result1,input) = parser1 input
- val (result2,input) = parser2 input
- in
- ((result1,result2),input)
- end;
-
-fun op >> (parser : 'a -> 'b * 'a, treatment) input =
- let
- val (result,input) = parser input
- in
- (treatment result, input)
- end;
-
-fun op >>++ (parser,treatment) input =
- let
- val (result,input) = parser input
- in
- treatment result input
- end;
-
-fun op || (parser1,parser2) input =
- parser1 input handle NoParse => parser2 input;
-
-fun first [] _ = raise NoParse
- | first (parser :: parsers) input = (parser || first parsers) input;
-
-fun mmany parser state input =
- let
- val (state,input) = parser state input
- in
- mmany parser state input
- end
- handle NoParse => (state,input);
-
-fun many parser =
- let
- fun sparser l = parser >> (fn x => x :: l)
- in
- mmany sparser [] >> rev
- end;
-
-fun atLeastOne p = (p ++ many p) >> op::;
-
-fun nothing input = ((),input);
-
-fun optional p = (p >> SOME) || (nothing >> K NONE);
-
-(* ------------------------------------------------------------------------- *)
-(* Stream-based. *)
-(* ------------------------------------------------------------------------- *)
-
-type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream
-
-fun everything parser =
- let
- fun f input =
- let
- val (result,input) = parser input
- in
- Stream.append (Stream.fromList result) (fn () => f input)
- end
- handle NoParse =>
- if Stream.null input then Stream.NIL else raise NoParse
- in
- f
- end;
-
-fun maybe p Stream.NIL = raise NoParse
- | maybe p (Stream.CONS (h,t)) =
- case p h of SOME r => (r, t ()) | NONE => raise NoParse;
-
-fun finished Stream.NIL = ((), Stream.NIL)
- | finished (Stream.CONS _) = raise NoParse;
-
-fun some p = maybe (fn x => if p x then SOME x else NONE);
-
-fun any input = some (K true) input;
-
-fun exact tok = some (fn item => item = tok);
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing and pretty-printing for infix operators. *)
-(* ------------------------------------------------------------------------- *)
-
-type infixities = {token : string, precedence : int, leftAssoc : bool} list;
-
-local
- fun unflatten ({token,precedence,leftAssoc}, ([],_)) =
- ([(leftAssoc, [token])], precedence)
- | unflatten ({token,precedence,leftAssoc}, ((a,l) :: dealt, p)) =
- if p = precedence then
- let
- val _ = leftAssoc = a orelse
- raise Bug "infix parser/printer: mixed assocs"
- in
- ((a, token :: l) :: dealt, p)
- end
- else
- ((leftAssoc,[token]) :: (a,l) :: dealt, precedence);
-in
- fun layerOps infixes =
- let
- val infixes = sortMap #precedence Int.compare infixes
- val (parsers,_) = foldl unflatten ([],0) infixes
- in
- parsers
- end;
-end;
-
-local
- fun chop (#" " :: chs) = let val (n,l) = chop chs in (n + 1, l) end
- | chop chs = (0,chs);
-
- fun nspaces n = funpow n (curry op^ " ") "";
-
- fun spacify tok =
- let
- val chs = explode tok
- val (r,chs) = chop (rev chs)
- val (l,chs) = chop (rev chs)
- in
- ((l,r), implode chs)
- end;
-
- fun lrspaces (l,r) =
- (if l = 0 then K () else C addString (nspaces l),
- if r = 0 then K () else C addBreak (r, 0));
-in
- fun opSpaces s = let val (l_r,s) = spacify s in (lrspaces l_r, s) end;
-
- val opClean = snd o spacify;
-end;
-
-val infixTokens : infixities -> string list =
- List.map (fn {token,...} => opClean token);
-
-fun parseGenInfix update sof toks parse inp =
- let
- val (e, rest) = parse inp
-
- val continue =
- case rest of
- Stream.NIL => NONE
- | Stream.CONS (h, t) => if mem h toks then SOME (h, t) else NONE
- in
- case continue of
- NONE => (sof e, rest)
- | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ())
- end;
-
-fun parseLeftInfix toks con =
- parseGenInfix (fn f => fn t => fn a => fn b => con (t, f a, b)) I toks;
-
-fun parseRightInfix toks con =
- parseGenInfix (fn f => fn t => fn a => fn b => f (con (t, a, b))) I toks;
-
-fun parseInfixes ops =
- let
- fun layeredOp (x,y) = (x, List.map opClean y)
-
- val layeredOps = List.map layeredOp (layerOps ops)
-
- fun iparser (a,t) = (if a then parseLeftInfix else parseRightInfix) t
-
- val iparsers = List.map iparser layeredOps
- in
- fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers
- end;
-
-fun ppGenInfix left toks =
- let
- val spc = List.map opSpaces toks
- in
- fn dest => fn ppSub =>
- let
- fun dest' tm =
- case dest tm of
- NONE => NONE
- | SOME (t, a, b) =>
- Option.map (pair (a,b)) (List.find (equal t o snd) spc)
-
- open PP
-
- fun ppGo pp (tmr as (tm,r)) =
- case dest' tm of
- NONE => ppSub pp tmr
- | SOME ((a,b),((lspc,rspc),tok)) =>
- ((if left then ppGo else ppSub) pp (a,true);
- lspc pp; addString pp tok; rspc pp;
- (if left then ppSub else ppGo) pp (b,r))
- in
- fn pp => fn tmr as (tm,_) =>
- case dest' tm of
- NONE => ppSub pp tmr
- | SOME _ => (beginBlock pp Inconsistent 0; ppGo pp tmr; endBlock pp)
- end
- end;
-
-fun ppLeftInfix toks = ppGenInfix true toks;
-
-fun ppRightInfix toks = ppGenInfix false toks;
-
-fun ppInfixes ops =
- let
- val layeredOps = layerOps ops
-
- val toks = List.concat (List.map (List.map opClean o snd) layeredOps)
-
- fun iprinter (a,t) = (if a then ppLeftInfix else ppRightInfix) t
-
- val iprinters = List.map iprinter layeredOps
- in
- fn dest => fn ppSub =>
- let
- fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters
-
- fun isOp t = case dest t of SOME (x,_,_) => mem x toks | _ => false
-
- open PP
-
- fun subpr pp (tmr as (tm,_)) =
- if isOp tm then
- (beginBlock pp Inconsistent 1; addString pp "(";
- printer subpr pp (tm, false); addString pp ")"; endBlock pp)
- else ppSub pp tmr
- in
- fn pp => fn tmr =>
- (beginBlock pp Inconsistent 0; printer subpr pp tmr; endBlock pp)
- end
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Quotations *)
-(* ------------------------------------------------------------------------- *)
-
-type 'a quotation = 'a frag list;
-
-fun parseQuotation printer parser quote =
- let
- fun expand (QUOTE q, s) = s ^ q
- | expand (ANTIQUOTE a, s) = s ^ printer a
-
- val string = foldl expand "" quote
- in
- parser string
- end;
-
-end
--- a/src/Tools/Metis/src/Portable.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-(* ========================================================================= *)
-(* ML SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Portable =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* The ML implementation. *)
-(* ------------------------------------------------------------------------- *)
-
-val ml : string
-
-(* ------------------------------------------------------------------------- *)
-(* Pointer equality using the run-time system. *)
-(* ------------------------------------------------------------------------- *)
-
-val pointerEqual : 'a * 'a -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
-(* ------------------------------------------------------------------------- *)
-
-val time : ('a -> 'b) -> 'a -> 'b
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-val CRITICAL: (unit -> 'a) -> 'a
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-val randomBool : unit -> bool
-
-val randomInt : int -> int (* n |-> [0,n) *)
-
-val randomReal : unit -> real (* () |-> [0,1] *)
-
-val randomWord : unit -> Word.word
-
-end
--- a/src/Tools/Metis/src/PortableIsabelle.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,47 +0,0 @@
-(* ========================================================================= *)
-(* Isabelle ML SPECIFIC FUNCTIONS *)
-(* ========================================================================= *)
-
-structure Portable :> Portable =
-struct
-
-(* ------------------------------------------------------------------------- *)
-(* The ML implementation. *)
-(* ------------------------------------------------------------------------- *)
-
-val ml = ml_system;
-
-(* ------------------------------------------------------------------------- *)
-(* Pointer equality using the run-time system. *)
-(* ------------------------------------------------------------------------- *)
-
-val pointerEqual = pointer_eq;
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications a la Mosml.time. *)
-(* ------------------------------------------------------------------------- *)
-
-val time = timeap;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-fun CRITICAL e = NAMED_CRITICAL "metis" e;
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-val randomWord = Random_Word.next_word;
-val randomBool = Random_Word.next_bool;
-fun randomInt n = Random_Word.next_int 0 (n - 1);
-val randomReal = Random_Word.next_real;
-
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Quotations a la Moscow ML. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a;
--- a/src/Tools/Metis/src/PortableMlton.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,92 +0,0 @@
-(* ========================================================================= *)
-(* MLTON SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Portable :> Portable =
-struct
-
-(* ------------------------------------------------------------------------- *)
-(* The ML implementation. *)
-(* ------------------------------------------------------------------------- *)
-
-val ml = "mlton";
-
-(* ------------------------------------------------------------------------- *)
-(* Pointer equality using the run-time system. *)
-(* ------------------------------------------------------------------------- *)
-
-val pointerEqual = MLton.eq;
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
-(* ------------------------------------------------------------------------- *)
-
-fun time f x =
- let
- fun p t =
- let
- val s = Time.fmt 3 t
- in
- case size (List.last (String.fields (fn x => x = #".") s)) of
- 3 => s
- | 2 => s ^ "0"
- | 1 => s ^ "00"
- | _ => raise Fail "Portable.time"
- end
-
- val c = Timer.startCPUTimer ()
-
- val r = Timer.startRealTimer ()
-
- fun pt () =
- let
- val {usr,sys} = Timer.checkCPUTimer c
- val real = Timer.checkRealTimer r
- in
- print
- ("User: " ^ p usr ^ " System: " ^ p sys ^
- " Real: " ^ p real ^ "\n")
- end
-
- val y = f x handle e => (pt (); raise e)
-
- val () = pt ()
- in
- y
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-fun CRITICAL e = e (); (*dummy*)
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-fun randomWord () = MLton.Random.rand ();
-
-fun randomBool () = Word.andb (randomWord (),0w1) = 0w0;
-
-fun randomInt 1 = 0
- | randomInt 2 = Word.toInt (Word.andb (randomWord (), 0w1))
- | randomInt 4 = Word.toInt (Word.andb (randomWord (), 0w3))
- | randomInt n = Word.toInt (Word.mod (randomWord (), Word.fromInt n));
-
-local
- fun wordToReal w = Real.fromInt (Word.toInt (Word.>> (w,0w1)))
-
- val normalizer = 1.0 / wordToReal (Word.notb 0w0);
-in
- fun randomReal () = normalizer * wordToReal (randomWord ());
-end;
-
-end
-
-(* ------------------------------------------------------------------------- *)
-(* Quotations a la Moscow ML. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a;
--- a/src/Tools/Metis/src/PortableMosml.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,120 +0,0 @@
-(* ========================================================================= *)
-(* MOSCOW ML SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Portable :> Portable =
-struct
-
-(* ------------------------------------------------------------------------- *)
-(* The ML implementation. *)
-(* ------------------------------------------------------------------------- *)
-
-val ml = "mosml";
-
-(* ------------------------------------------------------------------------- *)
-(* Pointer equality using the run-time system. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- val address : 'a -> int = Obj.magic
-in
- fun pointerEqual (x : 'a, y : 'a) = address x = address y
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
-(* ------------------------------------------------------------------------- *)
-
-val time = Mosml.time;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-fun CRITICAL e = e (); (*dummy*)
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- val gen = Random.newgenseed 1.0;
-in
- fun randomInt max = Random.range (0,max) gen;
-
- fun randomReal () = Random.random gen;
-end;
-
-fun randomBool () = randomInt 2 = 0;
-
-fun randomWord () =
- let
- val h = Word.fromInt (randomInt 65536)
- and l = Word.fromInt (randomInt 65536)
- in
- Word.orb (Word.<< (h,0w16), l)
- end;
-
-end
-
-(* ------------------------------------------------------------------------- *)
-(* Ensuring that interruptions (SIGINTs) are actually seen by the *)
-(* linked executable as Interrupt exceptions. *)
-(* ------------------------------------------------------------------------- *)
-
-prim_val catch_interrupt : bool -> unit = 1 "sys_catch_break";
-val _ = catch_interrupt true;
-
-(* ------------------------------------------------------------------------- *)
-(* Ad-hoc upgrading of the Moscow ML basis library. *)
-(* ------------------------------------------------------------------------- *)
-
-fun Array_copy {src,dst,di} =
- let
- open Array
- in
- copy {src = src, si = 0, len = NONE, dst = dst, di = di}
- end;
-
-fun Array_foldli f b v =
- let
- open Array
- in
- foldli f b (v,0,NONE)
- end;
-
-fun Array_foldri f b v =
- let
- open Array
- in
- foldri f b (v,0,NONE)
- end;
-
-fun Array_modifyi f a =
- let
- open Array
- in
- modifyi f (a,0,NONE)
- end;
-
-fun TextIO_inputLine h =
- let
- open TextIO
- in
- case inputLine h of "" => NONE | s => SOME s
- end;
-
-fun Vector_foldli f b v =
- let
- open Vector
- in
- foldli f b (v,0,NONE)
- end;
-
-fun Vector_mapi f v =
- let
- open Vector
- in
- mapi f (v,0,NONE)
- end;
--- a/src/Tools/Metis/src/Problem.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,57 +0,0 @@
-(* ========================================================================= *)
-(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Problem =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Problems. *)
-(* ------------------------------------------------------------------------- *)
-
-type problem = Thm.clause list
-
-val size : problem -> {clauses : int,
- literals : int,
- symbols : int,
- typedSymbols : int}
-
-val fromGoal : Formula.formula -> problem list
-
-val toClauses : problem -> Formula.formula
-
-val toString : problem -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Categorizing problems. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype propositional =
- Propositional
- | EffectivelyPropositional
- | NonPropositional
-
-datatype equality =
- NonEquality
- | Equality
- | PureEquality
-
-datatype horn =
- Trivial
- | Unit
- | DoubleHorn
- | Horn
- | NegativeHorn
- | NonHorn
-
-type category =
- {propositional : propositional,
- equality : equality,
- horn : horn}
-
-val categorize : problem -> category
-
-val categoryToString : category -> string
-
-end
--- a/src/Tools/Metis/src/Problem.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,185 +0,0 @@
-(* ========================================================================= *)
-(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Problem :> Problem =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Problems. *)
-(* ------------------------------------------------------------------------- *)
-
-type problem = Thm.clause list;
-
-fun size cls =
- {clauses = length cls,
- literals = foldl (fn (cl,n) => n + LiteralSet.size cl) 0 cls,
- symbols = foldl (fn (cl,n) => n + LiteralSet.symbols cl) 0 cls,
- typedSymbols = foldl (fn (cl,n) => n + LiteralSet.typedSymbols cl) 0 cls};
-
-fun checkFormula {models,checks} exp fm =
- let
- fun check 0 = true
- | check n =
- let
- val N = 3 + Portable.randomInt 3
- val M = Model.new {size = N, fixed = Model.fixedPure}
- val {T,F} = Model.checkFormula {maxChecks = checks} M fm
- in
- (if exp then F = 0 else T = 0) andalso check (n - 1)
- end
- in
- check models
- end;
-
-val checkGoal = checkFormula {models = 10, checks = 100} true;
-
-val checkClauses = checkFormula {models = 10, checks = 100} false;
-
-fun fromGoal goal =
- let
- fun fromLiterals fms = LiteralSet.fromList (map Literal.fromFormula fms)
-
- fun fromClause fm = fromLiterals (Formula.stripDisj fm)
-
- fun fromCnf fm = map fromClause (Formula.stripConj fm)
-
-(*DEBUG
- val () = if checkGoal goal then ()
- else raise Error "goal failed the finite model test"
-*)
-
- val refute = Formula.Not (Formula.generalize goal)
-
-(*TRACE2
- val () = Parser.ppTrace Formula.pp "Problem.fromGoal: refute" refute
-*)
-
- val cnfs = Normalize.cnf refute
-
-(*
- val () =
- let
- fun check fm =
- let
-(*TRACE2
- val () = Parser.ppTrace Formula.pp "Problem.fromGoal: cnf" fm
-*)
- in
- if checkClauses fm then ()
- else raise Error "cnf failed the finite model test"
- end
- in
- app check cnfs
- end
-*)
- in
- map fromCnf cnfs
- end;
-
-fun toClauses cls =
- let
- fun formulize cl =
- Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl)
- in
- Formula.listMkConj (map formulize cls)
- end;
-
-fun toString cls = Formula.toString (toClauses cls);
-
-(* ------------------------------------------------------------------------- *)
-(* Categorizing problems. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype propositional =
- Propositional
- | EffectivelyPropositional
- | NonPropositional;
-
-datatype equality =
- NonEquality
- | Equality
- | PureEquality;
-
-datatype horn =
- Trivial
- | Unit
- | DoubleHorn
- | Horn
- | NegativeHorn
- | NonHorn;
-
-type category =
- {propositional : propositional,
- equality : equality,
- horn : horn};
-
-fun categorize cls =
- let
- val rels =
- let
- fun f (cl,set) = NameAritySet.union set (LiteralSet.relations cl)
- in
- List.foldl f NameAritySet.empty cls
- end
-
- val funs =
- let
- fun f (cl,set) = NameAritySet.union set (LiteralSet.functions cl)
- in
- List.foldl f NameAritySet.empty cls
- end
-
- val propositional =
- if NameAritySet.allNullary rels then Propositional
- else if NameAritySet.allNullary funs then EffectivelyPropositional
- else NonPropositional
-
- val equality =
- if not (NameAritySet.member Atom.eqRelation rels) then NonEquality
- else if NameAritySet.size rels = 1 then PureEquality
- else Equality
-
- val horn =
- if List.exists LiteralSet.null cls then Trivial
- else if List.all (fn cl => LiteralSet.size cl = 1) cls then Unit
- else
- let
- fun pos cl = LiteralSet.count Literal.positive cl <= 1
- fun neg cl = LiteralSet.count Literal.negative cl <= 1
- in
- case (List.all pos cls, List.all neg cls) of
- (true,true) => DoubleHorn
- | (true,false) => Horn
- | (false,true) => NegativeHorn
- | (false,false) => NonHorn
- end
- in
- {propositional = propositional,
- equality = equality,
- horn = horn}
- end;
-
-fun categoryToString {propositional,equality,horn} =
- (case propositional of
- Propositional => "propositional"
- | EffectivelyPropositional => "effectively propositional"
- | NonPropositional => "non-propositional") ^
- ", " ^
- (case equality of
- NonEquality => "non-equality"
- | Equality => "equality"
- | PureEquality => "pure equality") ^
- ", " ^
- (case horn of
- Trivial => "trivial"
- | Unit => "unit"
- | DoubleHorn => "horn (and negative horn)"
- | Horn => "horn"
- | NegativeHorn => "negative horn"
- | NonHorn => "non-horn");
-
-end
--- a/src/Tools/Metis/src/Proof.sig Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,53 +0,0 @@
-(* ========================================================================= *)
-(* PROOFS IN FIRST ORDER LOGIC *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Proof =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* A type of first order logic proofs. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype inference =
- Axiom of LiteralSet.set
- | Assume of Atom.atom
- | Subst of Subst.subst * Thm.thm
- | Resolve of Atom.atom * Thm.thm * Thm.thm
- | Refl of Term.term
- | Equality of Literal.literal * Term.path * Term.term
-
-type proof = (Thm.thm * inference) list
-
-(* ------------------------------------------------------------------------- *)
-(* Reconstructing single inferences. *)
-(* ------------------------------------------------------------------------- *)
-
-val inferenceType : inference -> Thm.inferenceType
-
-val parents : inference -> Thm.thm list
-
-val inferenceToThm : inference -> Thm.thm
-
-val thmToInference : Thm.thm -> inference
-
-(* ------------------------------------------------------------------------- *)
-(* Reconstructing whole proofs. *)
-(* ------------------------------------------------------------------------- *)
-
-val proof : Thm.thm -> proof
-
-(* ------------------------------------------------------------------------- *)
-(* Printing. *)
-(* ------------------------------------------------------------------------- *)
-
-val ppInference : inference Parser.pp
-
-val inferenceToString : inference -> string
-
-val pp : proof Parser.pp
-
-val toString : proof -> string
-
-end
--- a/src/Tools/Metis/src/Proof.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,389 +0,0 @@
-(* ========================================================================= *)
-(* PROOFS IN FIRST ORDER LOGIC *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Proof :> Proof =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* A type of first order logic proofs. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype inference =
- Axiom of LiteralSet.set
- | Assume of Atom.atom
- | Subst of Subst.subst * Thm.thm
- | Resolve of Atom.atom * Thm.thm * Thm.thm
- | Refl of Term.term
- | Equality of Literal.literal * Term.path * Term.term;
-
-type proof = (Thm.thm * inference) list;
-
-(* ------------------------------------------------------------------------- *)
-(* Printing. *)
-(* ------------------------------------------------------------------------- *)
-
-fun inferenceType (Axiom _) = Thm.Axiom
- | inferenceType (Assume _) = Thm.Assume
- | inferenceType (Subst _) = Thm.Subst
- | inferenceType (Resolve _) = Thm.Resolve
- | inferenceType (Refl _) = Thm.Refl
- | inferenceType (Equality _) = Thm.Equality;
-
-local
- fun ppAssume pp atm = (Parser.addBreak pp (1,0); Atom.pp pp atm);
-
- fun ppSubst ppThm pp (sub,thm) =
- (Parser.addBreak pp (1,0);
- Parser.beginBlock pp Parser.Inconsistent 1;
- Parser.addString pp "{";
- Parser.ppBinop " =" Parser.ppString Subst.pp pp ("sub",sub);
- Parser.addString pp ",";
- Parser.addBreak pp (1,0);
- Parser.ppBinop " =" Parser.ppString ppThm pp ("thm",thm);
- Parser.addString pp "}";
- Parser.endBlock pp);
-
- fun ppResolve ppThm pp (res,pos,neg) =
- (Parser.addBreak pp (1,0);
- Parser.beginBlock pp Parser.Inconsistent 1;
- Parser.addString pp "{";
- Parser.ppBinop " =" Parser.ppString Atom.pp pp ("res",res);
- Parser.addString pp ",";
- Parser.addBreak pp (1,0);
- Parser.ppBinop " =" Parser.ppString ppThm pp ("pos",pos);
- Parser.addString pp ",";
- Parser.addBreak pp (1,0);
- Parser.ppBinop " =" Parser.ppString ppThm pp ("neg",neg);
- Parser.addString pp "}";
- Parser.endBlock pp);
-
- fun ppRefl pp tm = (Parser.addBreak pp (1,0); Term.pp pp tm);
-
- fun ppEquality pp (lit,path,res) =
- (Parser.addBreak pp (1,0);
- Parser.beginBlock pp Parser.Inconsistent 1;
- Parser.addString pp "{";
- Parser.ppBinop " =" Parser.ppString Literal.pp pp ("lit",lit);
- Parser.addString pp ",";
- Parser.addBreak pp (1,0);
- Parser.ppBinop " =" Parser.ppString Term.ppPath pp ("path",path);
- Parser.addString pp ",";
- Parser.addBreak pp (1,0);
- Parser.ppBinop " =" Parser.ppString Term.pp pp ("res",res);
- Parser.addString pp "}";
- Parser.endBlock pp);
-
- fun ppInf ppAxiom ppThm pp inf =
- let
- val infString = Thm.inferenceTypeToString (inferenceType inf)
- in
- Parser.beginBlock pp Parser.Inconsistent (size infString + 1);
- Parser.ppString pp infString;
- case inf of
- Axiom cl => ppAxiom pp cl
- | Assume x => ppAssume pp x
- | Subst x => ppSubst ppThm pp x
- | Resolve x => ppResolve ppThm pp x
- | Refl x => ppRefl pp x
- | Equality x => ppEquality pp x;
- Parser.endBlock pp
- end;
-
- fun ppAxiom pp cl =
- (Parser.addBreak pp (1,0);
- Parser.ppMap
- LiteralSet.toList
- (Parser.ppBracket "{" "}" (Parser.ppSequence "," Literal.pp)) pp cl);
-in
- val ppInference = ppInf ppAxiom Thm.pp;
-
- fun pp p prf =
- let
- fun thmString n = "(" ^ Int.toString n ^ ")"
-
- val prf = enumerate prf
-
- fun ppThm p th =
- let
- val cl = Thm.clause th
-
- fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl
- in
- case List.find pred prf of
- NONE => Parser.addString p "(?)"
- | SOME (n,_) => Parser.addString p (thmString n)
- end
-
- fun ppStep (n,(th,inf)) =
- let
- val s = thmString n
- in
- Parser.beginBlock p Parser.Consistent (1 + size s);
- Parser.addString p (s ^ " ");
- Thm.pp p th;
- Parser.addBreak p (2,0);
- Parser.ppBracket "[" "]" (ppInf (K (K ())) ppThm) p inf;
- Parser.endBlock p;
- Parser.addNewline p
- end
- in
- Parser.beginBlock p Parser.Consistent 0;
- Parser.addString p "START OF PROOF";
- Parser.addNewline p;
- app ppStep prf;
- Parser.addString p "END OF PROOF";
- Parser.addNewline p;
- Parser.endBlock p
- end
-(*DEBUG
- handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err);
-*)
-
-end;
-
-val inferenceToString = Parser.toString ppInference;
-
-val toString = Parser.toString pp;
-
-(* ------------------------------------------------------------------------- *)
-(* Reconstructing single inferences. *)
-(* ------------------------------------------------------------------------- *)
-
-fun parents (Axiom _) = []
- | parents (Assume _) = []
- | parents (Subst (_,th)) = [th]
- | parents (Resolve (_,th,th')) = [th,th']
- | parents (Refl _) = []
- | parents (Equality _) = [];
-
-fun inferenceToThm (Axiom cl) = Thm.axiom cl
- | inferenceToThm (Assume atm) = Thm.assume (true,atm)
- | inferenceToThm (Subst (sub,th)) = Thm.subst sub th
- | inferenceToThm (Resolve (atm,th,th')) = Thm.resolve (true,atm) th th'
- | inferenceToThm (Refl tm) = Thm.refl tm
- | inferenceToThm (Equality (lit,path,r)) = Thm.equality lit path r;
-
-local
- fun reconstructSubst cl cl' =
- let
- fun recon [] =
- let
-(*TRACE3
- val () = Parser.ppTrace LiteralSet.pp "reconstructSubst: cl" cl
- val () = Parser.ppTrace LiteralSet.pp "reconstructSubst: cl'" cl'
-*)
- in
- raise Bug "can't reconstruct Subst rule"
- end
- | recon (([],sub) :: others) =
- if LiteralSet.equal (LiteralSet.subst sub cl) cl' then sub
- else recon others
- | recon ((lit :: lits, sub) :: others) =
- let
- fun checkLit (lit',acc) =
- case total (Literal.match sub lit) lit' of
- NONE => acc
- | SOME sub => (lits,sub) :: acc
- in
- recon (LiteralSet.foldl checkLit others cl')
- end
- in
- Subst.normalize (recon [(LiteralSet.toList cl, Subst.empty)])
- end
-(*DEBUG
- handle Error err =>
- raise Bug ("Proof.recontructSubst: shouldn't fail:\n" ^ err);
-*)
-
- fun reconstructResolvant cl1 cl2 cl =
- (if not (LiteralSet.subset cl1 cl) then
- LiteralSet.pick (LiteralSet.difference cl1 cl)
- else if not (LiteralSet.subset cl2 cl) then
- Literal.negate (LiteralSet.pick (LiteralSet.difference cl2 cl))
- else
- (* A useless resolution, but we must reconstruct it anyway *)
- let
- val cl1' = LiteralSet.negate cl1
- and cl2' = LiteralSet.negate cl2
- val lits = LiteralSet.intersectList [cl1,cl1',cl2,cl2']
- in
- if not (LiteralSet.null lits) then LiteralSet.pick lits
- else raise Bug "can't reconstruct Resolve rule"
- end)
-(*DEBUG
- handle Error err =>
- raise Bug ("Proof.recontructResolvant: shouldn't fail:\n" ^ err);
-*)
-
- fun reconstructEquality cl =
- let
-(*TRACE3
- val () = Parser.ppTrace LiteralSet.pp "Proof.reconstructEquality: cl" cl
-*)
-
- fun sync s t path (f,a) (f',a') =
- if f <> f' orelse length a <> length a' then NONE
- else
- case List.filter (op<> o snd) (enumerate (zip a a')) of
- [(i,(tm,tm'))] =>
- let
- val path = i :: path
- in
- if tm = s andalso tm' = t then SOME (rev path)
- else
- case (tm,tm') of
- (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a'
- | _ => NONE
- end
- | _ => NONE
-
- fun recon (neq,(pol,atm),(pol',atm')) =
- if pol = pol' then NONE
- else
- let
- val (s,t) = Literal.destNeq neq
-
- val path =
- if s <> t then sync s t [] atm atm'
- else if atm <> atm' then NONE
- else Atom.find (equal s) atm
- in
- case path of
- SOME path => SOME ((pol',atm),path,t)
- | NONE => NONE
- end
-
- val candidates =
- case List.partition Literal.isNeq (LiteralSet.toList cl) of
- ([l1],[l2,l3]) => [(l1,l2,l3),(l1,l3,l2)]
- | ([l1,l2],[l3]) => [(l1,l2,l3),(l1,l3,l2),(l2,l1,l3),(l2,l3,l1)]
- | ([l1],[l2]) => [(l1,l1,l2),(l1,l2,l1)]
- | _ => raise Bug "reconstructEquality: malformed"
-
-(*TRACE3
- val ppCands =
- Parser.ppList (Parser.ppTriple Literal.pp Literal.pp Literal.pp)
- val () = Parser.ppTrace ppCands
- "Proof.reconstructEquality: candidates" candidates
-*)
- in
- case first recon candidates of
- SOME info => info
- | NONE => raise Bug "can't reconstruct Equality rule"
- end
-(*DEBUG
- handle Error err =>
- raise Bug ("Proof.recontructEquality: shouldn't fail:\n" ^ err);
-*)
-
- fun reconstruct cl (Thm.Axiom,[]) = Axiom cl
- | reconstruct cl (Thm.Assume,[]) =
- (case LiteralSet.findl Literal.positive cl of
- SOME (_,atm) => Assume atm
- | NONE => raise Bug "malformed Assume inference")
- | reconstruct cl (Thm.Subst,[th]) =
- Subst (reconstructSubst (Thm.clause th) cl, th)
- | reconstruct cl (Thm.Resolve,[th1,th2]) =
- let
- val cl1 = Thm.clause th1
- and cl2 = Thm.clause th2
- val (pol,atm) = reconstructResolvant cl1 cl2 cl
- in
- if pol then Resolve (atm,th1,th2) else Resolve (atm,th2,th1)
- end
- | reconstruct cl (Thm.Refl,[]) =
- (case LiteralSet.findl (K true) cl of
- SOME lit => Refl (Literal.destRefl lit)
- | NONE => raise Bug "malformed Refl inference")
- | reconstruct cl (Thm.Equality,[]) = Equality (reconstructEquality cl)
- | reconstruct _ _ = raise Bug "malformed inference";
-in
- fun thmToInference th =
- let
-(*TRACE3
- val () = Parser.ppTrace Thm.pp "Proof.thmToInference: th" th
-*)
-
- val cl = Thm.clause th
-
- val thmInf = Thm.inference th
-
-(*TRACE3
- val ppThmInf = Parser.ppPair Thm.ppInferenceType (Parser.ppList Thm.pp)
- val () = Parser.ppTrace ppThmInf "Proof.thmToInference: thmInf" thmInf
-*)
-
- val inf = reconstruct cl thmInf
-
-(*TRACE3
- val () = Parser.ppTrace ppInference "Proof.thmToInference: inf" inf
-*)
-(*DEBUG
- val () =
- let
- val th' = inferenceToThm inf
- in
- if LiteralSet.equal (Thm.clause th') cl then ()
- else
- raise
- Bug
- ("Proof.thmToInference: bad inference reconstruction:" ^
- "\n th = " ^ Thm.toString th ^
- "\n inf = " ^ inferenceToString inf ^
- "\n inf th = " ^ Thm.toString th')
- end
-*)
- in
- inf
- end
-(*DEBUG
- handle Error err =>
- raise Bug ("Proof.thmToInference: shouldn't fail:\n" ^ err);
-*)
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Reconstructing whole proofs. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun thmCompare (th1,th2) =
- LiteralSet.compare (Thm.clause th1, Thm.clause th2);
-
- fun buildProof (th,(m,l)) =
- if Map.inDomain th m then (m,l)
- else
- let
- val (_,deps) = Thm.inference th
- val (m,l) = foldl buildProof (m,l) deps
- in
- if Map.inDomain th m then (m,l)
- else
- let
- val l = (th, thmToInference th) :: l
- in
- (Map.insert m (th,l), l)
- end
- end;
-in
- fun proof th =
- let
-(*TRACE3
- val () = Parser.ppTrace Thm.pp "Proof.proof: th" th
-*)
- val (m,_) = buildProof (th, (Map.new thmCompare, []))
-(*TRACE3
- val () = Parser.ppTrace Parser.ppInt "Proof.proof: size" (Map.size m)
-*)
- in
- case Map.peek m th of
- SOME l => rev l
- | NONE => raise Bug "Proof.proof"
- end;
-end;
-
-end
--- a/src/Tools/Metis/src/RandomMap.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,623 +0,0 @@
-(* ========================================================================= *)
-(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure RandomMap :> Map =
-struct
-
-exception Bug = Useful.Bug;
-
-exception Error = Useful.Error;
-
-val pointerEqual = Portable.pointerEqual;
-
-val K = Useful.K;
-
-val snd = Useful.snd;
-
-val randomInt = Portable.randomInt;
-
-val randomWord = Portable.randomWord;
-
-(* ------------------------------------------------------------------------- *)
-(* Random search trees. *)
-(* ------------------------------------------------------------------------- *)
-
-type priority = Word.word;
-
-datatype ('a,'b) tree =
- E
- | T of
- {size : int,
- priority : priority,
- left : ('a,'b) tree,
- key : 'a,
- value : 'b,
- right : ('a,'b) tree};
-
-type ('a,'b) node =
- {size : int,
- priority : priority,
- left : ('a,'b) tree,
- key : 'a,
- value : 'b,
- right : ('a,'b) tree};
-
-datatype ('a,'b) map = Map of ('a * 'a -> order) * ('a,'b) tree;
-
-(* ------------------------------------------------------------------------- *)
-(* Random priorities. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- val randomPriority = randomWord;
-
- val priorityOrder = Word.compare;
-in
- fun treeSingleton (key,value) =
- T {size = 1, priority = randomPriority (),
- left = E, key = key, value = value, right = E};
-
- fun nodePriorityOrder cmp (x1 : ('a,'b) node, x2 : ('a,'b) node) =
- let
- val {priority = p1, key = k1, ...} = x1
- and {priority = p2, key = k2, ...} = x2
- in
- case priorityOrder (p1,p2) of
- LESS => LESS
- | EQUAL => cmp (k1,k2)
- | GREATER => GREATER
- end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Debugging functions. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun checkSizes E = 0
- | checkSizes (T {size,left,right,...}) =
- let
- val l = checkSizes left
- and r = checkSizes right
- val () = if l + 1 + r = size then () else raise Error "wrong size"
- in
- size
- end;
-
- fun checkSorted _ x E = x
- | checkSorted cmp x (T {left,key,right,...}) =
- let
- val x = checkSorted cmp x left
- val () =
- case x of
- NONE => ()
- | SOME k =>
- case cmp (k,key) of
- LESS => ()
- | EQUAL => raise Error "duplicate keys"
- | GREATER => raise Error "unsorted"
- in
- checkSorted cmp (SOME key) right
- end;
-
- fun checkPriorities _ E = NONE
- | checkPriorities cmp (T (x as {left,right,...})) =
- let
- val () =
- case checkPriorities cmp left of
- NONE => ()
- | SOME l =>
- case nodePriorityOrder cmp (l,x) of
- LESS => ()
- | EQUAL => raise Error "left child has equal key"
- | GREATER => raise Error "left child has greater priority"
- val () =
- case checkPriorities cmp right of
- NONE => ()
- | SOME r =>
- case nodePriorityOrder cmp (r,x) of
- LESS => ()
- | EQUAL => raise Error "right child has equal key"
- | GREATER => raise Error "right child has greater priority"
- in
- SOME x
- end;
-in
- fun checkWellformed s (m as Map (cmp,tree)) =
- (let
- val _ = checkSizes tree
- val _ = checkSorted cmp NONE tree
- val _ = checkPriorities cmp tree
- in
- m
- end
- handle Error err => raise Bug err)
- handle Bug bug => raise Bug (s ^ "\nRandomMap.checkWellformed: " ^ bug);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun comparison (Map (cmp,_)) = cmp;
-
-fun new cmp = Map (cmp,E);
-
-fun treeSize E = 0
- | treeSize (T {size = s, ...}) = s;
-
-fun size (Map (_,tree)) = treeSize tree;
-
-fun mkT p l k v r =
- T {size = treeSize l + 1 + treeSize r, priority = p,
- left = l, key = k, value = v, right = r};
-
-fun singleton cmp key_value = Map (cmp, treeSingleton key_value);
-
-local
- fun treePeek cmp E pkey = NONE
- | treePeek cmp (T {left,key,value,right,...}) pkey =
- case cmp (pkey,key) of
- LESS => treePeek cmp left pkey
- | EQUAL => SOME value
- | GREATER => treePeek cmp right pkey
-in
- fun peek (Map (cmp,tree)) key = treePeek cmp tree key;
-end;
-
-(* treeAppend assumes that every element of the first tree is less than *)
-(* every element of the second tree. *)
-
-fun treeAppend _ t1 E = t1
- | treeAppend _ E t2 = t2
- | treeAppend cmp (t1 as T x1) (t2 as T x2) =
- case nodePriorityOrder cmp (x1,x2) of
- LESS =>
- let
- val {priority = p2,
- left = l2, key = k2, value = v2, right = r2, ...} = x2
- in
- mkT p2 (treeAppend cmp t1 l2) k2 v2 r2
- end
- | EQUAL => raise Bug "RandomSet.treeAppend: equal keys"
- | GREATER =>
- let
- val {priority = p1,
- left = l1, key = k1, value = v1, right = r1, ...} = x1
- in
- mkT p1 l1 k1 v1 (treeAppend cmp r1 t2)
- end;
-
-(* nodePartition splits the node into three parts: the keys comparing less *)
-(* than the supplied key, an optional equal key, and the keys comparing *)
-(* greater. *)
-
-local
- fun mkLeft [] t = t
- | mkLeft (({priority,left,key,value,...} : ('a,'b) node) :: xs) t =
- mkLeft xs (mkT priority left key value t);
-
- fun mkRight [] t = t
- | mkRight (({priority,key,value,right,...} : ('a,'b) node) :: xs) t =
- mkRight xs (mkT priority t key value right);
-
- fun treePart _ _ lefts rights E = (mkLeft lefts E, NONE, mkRight rights E)
- | treePart cmp pkey lefts rights (T x) = nodePart cmp pkey lefts rights x
- and nodePart cmp pkey lefts rights (x as {left,key,value,right,...}) =
- case cmp (pkey,key) of
- LESS => treePart cmp pkey lefts (x :: rights) left
- | EQUAL => (mkLeft lefts left, SOME (key,value), mkRight rights right)
- | GREATER => treePart cmp pkey (x :: lefts) rights right;
-in
- fun nodePartition cmp x pkey = nodePart cmp pkey [] [] x;
-end;
-
-(* union first calls treeCombineRemove, to combine the values *)
-(* for equal keys into the first map and remove them from the second map. *)
-(* Note that the combined key is always the one from the second map. *)
-
-local
- fun treeCombineRemove _ _ t1 E = (t1,E)
- | treeCombineRemove _ _ E t2 = (E,t2)
- | treeCombineRemove cmp f (t1 as T x1) (t2 as T x2) =
- let
- val {priority = p1,
- left = l1, key = k1, value = v1, right = r1, ...} = x1
- val (l2,k2_v2,r2) = nodePartition cmp x2 k1
- val (l1,l2) = treeCombineRemove cmp f l1 l2
- and (r1,r2) = treeCombineRemove cmp f r1 r2
- in
- case k2_v2 of
- NONE =>
- if treeSize l2 + treeSize r2 = #size x2 then (t1,t2)
- else (mkT p1 l1 k1 v1 r1, treeAppend cmp l2 r2)
- | SOME (k2,v2) =>
- case f (v1,v2) of
- NONE => (treeAppend cmp l1 r1, treeAppend cmp l2 r2)
- | SOME v => (mkT p1 l1 k2 v r1, treeAppend cmp l2 r2)
- end;
-
- fun treeUnionDisjoint _ t1 E = t1
- | treeUnionDisjoint _ E t2 = t2
- | treeUnionDisjoint cmp (T x1) (T x2) =
- case nodePriorityOrder cmp (x1,x2) of
- LESS => nodeUnionDisjoint cmp x2 x1
- | EQUAL => raise Bug "RandomSet.unionDisjoint: equal keys"
- | GREATER => nodeUnionDisjoint cmp x1 x2
- and nodeUnionDisjoint cmp x1 x2 =
- let
- val {priority = p1,
- left = l1, key = k1, value = v1, right = r1, ...} = x1
- val (l2,_,r2) = nodePartition cmp x2 k1
- val l = treeUnionDisjoint cmp l1 l2
- and r = treeUnionDisjoint cmp r1 r2
- in
- mkT p1 l k1 v1 r
- end;
-in
- fun union f (m1 as Map (cmp,t1)) (Map (_,t2)) =
- if pointerEqual (t1,t2) then m1
- else
- let
- val (t1,t2) = treeCombineRemove cmp f t1 t2
- in
- Map (cmp, treeUnionDisjoint cmp t1 t2)
- end;
-end;
-
-(*DEBUG
-val union = fn f => fn t1 => fn t2 =>
- checkWellformed "RandomMap.union: result"
- (union f (checkWellformed "RandomMap.union: input 1" t1)
- (checkWellformed "RandomMap.union: input 2" t2));
-*)
-
-(* intersect is a simple case of the union algorithm. *)
-
-local
- fun treeIntersect _ _ _ E = E
- | treeIntersect _ _ E _ = E
- | treeIntersect cmp f (t1 as T x1) (t2 as T x2) =
- let
- val {priority = p1,
- left = l1, key = k1, value = v1, right = r1, ...} = x1
- val (l2,k2_v2,r2) = nodePartition cmp x2 k1
- val l = treeIntersect cmp f l1 l2
- and r = treeIntersect cmp f r1 r2
- in
- case k2_v2 of
- NONE => treeAppend cmp l r
- | SOME (k2,v2) =>
- case f (v1,v2) of
- NONE => treeAppend cmp l r
- | SOME v => mkT p1 l k2 v r
- end;
-in
- fun intersect f (m1 as Map (cmp,t1)) (Map (_,t2)) =
- if pointerEqual (t1,t2) then m1
- else Map (cmp, treeIntersect cmp f t1 t2);
-end;
-
-(*DEBUG
-val intersect = fn f => fn t1 => fn t2 =>
- checkWellformed "RandomMap.intersect: result"
- (intersect f (checkWellformed "RandomMap.intersect: input 1" t1)
- (checkWellformed "RandomMap.intersect: input 2" t2));
-*)
-
-(* delete raises an exception if the supplied key is not found, which *)
-(* makes it simpler to maximize sharing. *)
-
-local
- fun treeDelete _ E _ = raise Error "RandomMap.delete: element not found"
- | treeDelete cmp (T {priority,left,key,value,right,...}) dkey =
- case cmp (dkey,key) of
- LESS => mkT priority (treeDelete cmp left dkey) key value right
- | EQUAL => treeAppend cmp left right
- | GREATER => mkT priority left key value (treeDelete cmp right dkey);
-in
- fun delete (Map (cmp,tree)) key = Map (cmp, treeDelete cmp tree key);
-end;
-
-(*DEBUG
-val delete = fn t => fn x =>
- checkWellformed "RandomMap.delete: result"
- (delete (checkWellformed "RandomMap.delete: input" t) x);
-*)
-
-(* Set difference on domains *)
-
-local
- fun treeDifference _ t1 E = t1
- | treeDifference _ E _ = E
- | treeDifference cmp (t1 as T x1) (T x2) =
- let
- val {size = s1, priority = p1,
- left = l1, key = k1, value = v1, right = r1} = x1
- val (l2,k2_v2,r2) = nodePartition cmp x2 k1
- val l = treeDifference cmp l1 l2
- and r = treeDifference cmp r1 r2
- in
- if Option.isSome k2_v2 then treeAppend cmp l r
- else if treeSize l + treeSize r + 1 = s1 then t1
- else mkT p1 l k1 v1 r
- end;
-in
- fun difference (Map (cmp,tree1)) (Map (_,tree2)) =
- Map (cmp, treeDifference cmp tree1 tree2);
-end;
-
-(*DEBUG
-val difference = fn t1 => fn t2 =>
- checkWellformed "RandomMap.difference: result"
- (difference (checkWellformed "RandomMap.difference: input 1" t1)
- (checkWellformed "RandomMap.difference: input 2" t2));
-*)
-
-(* subsetDomain is mainly used when using maps as sets. *)
-
-local
- fun treeSubsetDomain _ E _ = true
- | treeSubsetDomain _ _ E = false
- | treeSubsetDomain cmp (t1 as T x1) (T x2) =
- let
- val {size = s1, left = l1, key = k1, right = r1, ...} = x1
- and {size = s2, ...} = x2
- in
- s1 <= s2 andalso
- let
- val (l2,k2_v2,r2) = nodePartition cmp x2 k1
- in
- Option.isSome k2_v2 andalso
- treeSubsetDomain cmp l1 l2 andalso
- treeSubsetDomain cmp r1 r2
- end
- end;
-in
- fun subsetDomain (Map (cmp,tree1)) (Map (_,tree2)) =
- pointerEqual (tree1,tree2) orelse
- treeSubsetDomain cmp tree1 tree2;
-end;
-
-(* Map equality *)
-
-local
- fun treeEqual _ _ E E = true
- | treeEqual _ _ E _ = false
- | treeEqual _ _ _ E = false
- | treeEqual cmp veq (t1 as T x1) (T x2) =
- let
- val {size = s1, left = l1, key = k1, value = v1, right = r1, ...} = x1
- and {size = s2, ...} = x2
- in
- s1 = s2 andalso
- let
- val (l2,k2_v2,r2) = nodePartition cmp x2 k1
- in
- (case k2_v2 of NONE => false | SOME (_,v2) => veq v1 v2) andalso
- treeEqual cmp veq l1 l2 andalso
- treeEqual cmp veq r1 r2
- end
- end;
-in
- fun equal veq (Map (cmp,tree1)) (Map (_,tree2)) =
- pointerEqual (tree1,tree2) orelse
- treeEqual cmp veq tree1 tree2;
-end;
-
-(* mapPartial is the basic function for preserving the tree structure. *)
-(* It applies the argument function to the elements *in order*. *)
-
-local
- fun treeMapPartial cmp _ E = E
- | treeMapPartial cmp f (T {priority,left,key,value,right,...}) =
- let
- val left = treeMapPartial cmp f left
- and value' = f (key,value)
- and right = treeMapPartial cmp f right
- in
- case value' of
- NONE => treeAppend cmp left right
- | SOME value => mkT priority left key value right
- end;
-in
- fun mapPartial f (Map (cmp,tree)) = Map (cmp, treeMapPartial cmp f tree);
-end;
-
-(* map is a primitive function for efficiency reasons. *)
-(* It also applies the argument function to the elements *in order*. *)
-
-local
- fun treeMap _ E = E
- | treeMap f (T {size,priority,left,key,value,right}) =
- let
- val left = treeMap f left
- and value = f (key,value)
- and right = treeMap f right
- in
- T {size = size, priority = priority, left = left,
- key = key, value = value, right = right}
- end;
-in
- fun map f (Map (cmp,tree)) = Map (cmp, treeMap f tree);
-end;
-
-(* nth picks the nth smallest key/value (counting from 0). *)
-
-local
- fun treeNth E _ = raise Subscript
- | treeNth (T {left,key,value,right,...}) n =
- let
- val k = treeSize left
- in
- if n = k then (key,value)
- else if n < k then treeNth left n
- else treeNth right (n - (k + 1))
- end;
-in
- fun nth (Map (_,tree)) n = treeNth tree n;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators. *)
-(* ------------------------------------------------------------------------- *)
-
-fun leftSpine E acc = acc
- | leftSpine (t as T {left,...}) acc = leftSpine left (t :: acc);
-
-fun rightSpine E acc = acc
- | rightSpine (t as T {right,...}) acc = rightSpine right (t :: acc);
-
-datatype ('key,'a) iterator =
- LR of ('key * 'a) * ('key,'a) tree * ('key,'a) tree list
- | RL of ('key * 'a) * ('key,'a) tree * ('key,'a) tree list;
-
-fun mkLR [] = NONE
- | mkLR (T {key,value,right,...} :: l) = SOME (LR ((key,value),right,l))
- | mkLR (E :: _) = raise Bug "RandomMap.mkLR";
-
-fun mkRL [] = NONE
- | mkRL (T {key,value,left,...} :: l) = SOME (RL ((key,value),left,l))
- | mkRL (E :: _) = raise Bug "RandomMap.mkRL";
-
-fun mkIterator (Map (_,tree)) = mkLR (leftSpine tree []);
-
-fun mkRevIterator (Map (_,tree)) = mkRL (rightSpine tree []);
-
-fun readIterator (LR (key_value,_,_)) = key_value
- | readIterator (RL (key_value,_,_)) = key_value;
-
-fun advanceIterator (LR (_,next,l)) = mkLR (leftSpine next l)
- | advanceIterator (RL (_,next,l)) = mkRL (rightSpine next l);
-
-(* ------------------------------------------------------------------------- *)
-(* Derived operations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun null m = size m = 0;
-
-fun get m key =
- case peek m key of
- NONE => raise Error "RandomMap.get: element not found"
- | SOME value => value;
-
-fun inDomain key m = Option.isSome (peek m key);
-
-fun insert m key_value =
- union (SOME o snd) m (singleton (comparison m) key_value);
-
-(*DEBUG
-val insert = fn m => fn x =>
- checkWellformed "RandomMap.insert: result"
- (insert (checkWellformed "RandomMap.insert: input" m) x);
-*)
-
-local
- fun fold _ NONE acc = acc
- | fold f (SOME iter) acc =
- let
- val (key,value) = readIterator iter
- in
- fold f (advanceIterator iter) (f (key,value,acc))
- end;
-in
- fun foldl f b m = fold f (mkIterator m) b;
-
- fun foldr f b m = fold f (mkRevIterator m) b;
-end;
-
-local
- fun find _ NONE = NONE
- | find pred (SOME iter) =
- let
- val key_value = readIterator iter
- in
- if pred key_value then SOME key_value
- else find pred (advanceIterator iter)
- end;
-in
- fun findl p m = find p (mkIterator m);
-
- fun findr p m = find p (mkRevIterator m);
-end;
-
-local
- fun first _ NONE = NONE
- | first f (SOME iter) =
- let
- val key_value = readIterator iter
- in
- case f key_value of
- NONE => first f (advanceIterator iter)
- | s => s
- end;
-in
- fun firstl f m = first f (mkIterator m);
-
- fun firstr f m = first f (mkRevIterator m);
-end;
-
-fun fromList cmp l = List.foldl (fn (k_v,m) => insert m k_v) (new cmp) l;
-
-fun insertList m l = union (SOME o snd) m (fromList (comparison m) l);
-
-fun filter p =
- let
- fun f (key_value as (_,value)) =
- if p key_value then SOME value else NONE
- in
- mapPartial f
- end;
-
-fun app f m = foldl (fn (key,value,()) => f (key,value)) () m;
-
-fun transform f = map (fn (_,value) => f value);
-
-fun toList m = foldr (fn (key,value,l) => (key,value) :: l) [] m;
-
-fun domain m = foldr (fn (key,_,l) => key :: l) [] m;
-
-fun exists p m = Option.isSome (findl p m);
-
-fun all p m = not (exists (not o p) m);
-
-fun random m =
- case size m of
- 0 => raise Error "RandomMap.random: empty"
- | n => nth m (randomInt n);
-
-local
- fun iterCompare _ _ NONE NONE = EQUAL
- | iterCompare _ _ NONE (SOME _) = LESS
- | iterCompare _ _ (SOME _) NONE = GREATER
- | iterCompare kcmp vcmp (SOME i1) (SOME i2) =
- keyIterCompare kcmp vcmp (readIterator i1) (readIterator i2) i1 i2
-
- and keyIterCompare kcmp vcmp (k1,v1) (k2,v2) i1 i2 =
- case kcmp (k1,k2) of
- LESS => LESS
- | EQUAL =>
- (case vcmp (v1,v2) of
- LESS => LESS
- | EQUAL =>
- iterCompare kcmp vcmp (advanceIterator i1) (advanceIterator i2)
- | GREATER => GREATER)
- | GREATER => GREATER;
-in
- fun compare vcmp (m1,m2) =
- if pointerEqual (m1,m2) then EQUAL
- else
- case Int.compare (size m1, size m2) of
- LESS => LESS
- | EQUAL =>
- iterCompare (comparison m1) vcmp (mkIterator m1) (mkIterator m2)
- | GREATER => GREATER;
-end;
-
-fun equalDomain m1 m2 = equal (K (K true)) m1 m2;
-
-fun toString m = "<" ^ (if null m then "" else Int.toString (size m)) ^ ">";
-
-end
--- a/src/Tools/Metis/src/RandomSet.sml Mon Sep 13 20:27:40 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,607 +0,0 @@
-(* ========================================================================= *)
-(* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure RandomSet :> Set =
-struct
-
-exception Bug = Useful.Bug;
-
-exception Error = Useful.Error;
-
-val pointerEqual = Portable.pointerEqual;
-
-val K = Useful.K;
-
-val snd = Useful.snd;
-
-val randomInt = Portable.randomInt;
-
-val randomWord = Portable.randomWord;
-
-(* ------------------------------------------------------------------------- *)
-(* Random search trees. *)
-(* ------------------------------------------------------------------------- *)
-
-type priority = Word.word;
-
-datatype 'a tree =
- E
- | T of
- {size : int,
- priority : priority,
- left : 'a tree,
- key : 'a,
- right : 'a tree};
-
-type 'a node =
- {size : int,
- priority : priority,
- left : 'a tree,
- key : 'a,
- right : 'a tree};
-
-datatype 'a set = Set of ('a * 'a -> order) * 'a tree;
-
-(* ------------------------------------------------------------------------- *)
-(* Random priorities. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- val randomPriority = randomWord;
-
- val priorityOrder = Word.compare;
-in
- fun treeSingleton key =
- T {size = 1, priority = randomPriority (),
- left = E, key = key, right = E};
-
- fun nodePriorityOrder cmp (x1 : 'a node, x2 : 'a node) =
- let
- val {priority = p1, key = k1, ...} = x1
- and {priority = p2, key = k2, ...} = x2
- in
- case priorityOrder (p1,p2) of
- LESS => LESS
- | EQUAL => cmp (k1,k2)
- | GREATER => GREATER
- end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Debugging functions. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun checkSizes E = 0
- | checkSizes (T {size,left,right,...}) =
- let
- val l = checkSizes left
- and r = checkSizes right
- val () = if l + 1 + r = size then () else raise Error "wrong size"
- in
- size
- end
-
- fun checkSorted _ x E = x
- | checkSorted cmp x (T {left,key,right,...}) =
- let
- val x = checkSorted cmp x left
- val () =
- case x of
- NONE => ()
- | SOME k =>
- case cmp (k,key) of
- LESS => ()
- | EQUAL => raise Error "duplicate keys"
- | GREATER => raise Error "unsorted"
- in
- checkSorted cmp (SOME key) right
- end;
-
- fun checkPriorities _ E = NONE
- | checkPriorities cmp (T (x as {left,right,...})) =
- let
- val () =
- case checkPriorities cmp left of
- NONE => ()
- | SOME l =>
- case nodePriorityOrder cmp (l,x) of
- LESS => ()
- | EQUAL => raise Error "left child has equal key"
- | GREATER => raise Error "left child has greater priority"
- val () =
- case checkPriorities cmp right of
- NONE => ()
- | SOME r =>
- case nodePriorityOrder cmp (r,x) of
- LESS => ()
- | EQUAL => raise Error "right child has equal key"
- | GREATER => raise Error "right child has greater priority"
- in
- SOME x
- end;
-in
- fun checkWellformed s (set as Set (cmp,tree)) =
- (let
- val _ = checkSizes tree
- val _ = checkSorted cmp NONE tree
- val _ = checkPriorities cmp tree
- in
- set
- end
- handle Error err => raise Bug err)
- handle Bug bug => raise Bug (s ^ "\nRandomSet.checkWellformed: " ^ bug);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun comparison (Set (cmp,_)) = cmp;
-
-fun empty cmp = Set (cmp,E);
-
-fun treeSize E = 0
- | treeSize (T {size = s, ...}) = s;
-
-fun size (Set (_,tree)) = treeSize tree;
-
-fun mkT p l k r =
- T {size = treeSize l + 1 + treeSize r, priority = p,
- left = l, key = k, right = r};
-
-fun singleton cmp key = Set (cmp, treeSingleton key);
-
-local
- fun treePeek cmp E pkey = NONE
- | treePeek cmp (T {left,key,right,...}) pkey =
- case cmp (pkey,key) of
- LESS => treePeek cmp left pkey
- | EQUAL => SOME key
- | GREATER => treePeek cmp right pkey
-in
- fun peek (Set (cmp,tree)) key = treePeek cmp tree key;
-end;
-
-(* treeAppend assumes that every element of the first tree is less than *)
-(* every element of the second tree. *)
-
-fun treeAppend _ t1 E = t1
- | treeAppend _ E t2 = t2
- | treeAppend cmp (t1 as T x1) (t2 as T x2) =
- case nodePriorityOrder cmp (x1,x2) of
- LESS =>
- let
- val {priority = p2, left = l2, key = k2, right = r2, ...} = x2
- in
- mkT p2 (treeAppend cmp t1 l2) k2 r2
- end
- | EQUAL => raise Bug "RandomSet.treeAppend: equal keys"
- | GREATER =>
- let
- val {priority = p1, left = l1, key = k1, right = r1, ...} = x1
-