diff --git a/src/HOL/Nitpick_Examples/minipick.ML b/src/HOL/Nitpick_Examples/minipick.ML --- a/src/HOL/Nitpick_Examples/minipick.ML +++ b/src/HOL/Nitpick_Examples/minipick.ML @@ -1,470 +1,469 @@ (* Title: HOL/Nitpick_Examples/minipick.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2009-2010 Finite model generation for HOL formulas using Kodkod, minimalistic version. *) signature MINIPICK = sig val minipick : Proof.context -> int -> term -> string val minipick_expect : Proof.context -> string -> int -> term -> unit end; structure Minipick : MINIPICK = struct open Kodkod open Nitpick_Util open Nitpick_HOL open Nitpick_Peephole open Nitpick_Kodkod datatype rep = S_Rep | R_Rep of bool fun check_type ctxt raw_infinite (Type (\<^type_name>\fun\, Ts)) = List.app (check_type ctxt raw_infinite) Ts | check_type ctxt raw_infinite (Type (\<^type_name>\prod\, Ts)) = List.app (check_type ctxt raw_infinite) Ts | check_type _ _ \<^typ>\bool\ = () | check_type _ _ (TFree (_, \<^sort>\{}\)) = () | check_type _ _ (TFree (_, \<^sort>\HOL.type\)) = () | check_type ctxt raw_infinite T = if raw_infinite T then () else error ("Not supported: Type " ^ quote (Syntax.string_of_typ ctxt T) ^ ".") fun atom_schema_of S_Rep card (Type (\<^type_name>\fun\, [T1, T2])) = replicate_list (card T1) (atom_schema_of S_Rep card T2) | atom_schema_of (R_Rep true) card (Type (\<^type_name>\fun\, [T1, \<^typ>\bool\])) = atom_schema_of S_Rep card T1 | atom_schema_of (rep as R_Rep _) card (Type (\<^type_name>\fun\, [T1, T2])) = atom_schema_of S_Rep card T1 @ atom_schema_of rep card T2 | atom_schema_of _ card (Type (\<^type_name>\prod\, Ts)) = maps (atom_schema_of S_Rep card) Ts | atom_schema_of _ card T = [card T] val arity_of = length ooo atom_schema_of val atom_seqs_of = map (AtomSeq o rpair 0) ooo atom_schema_of val atom_seq_product_of = foldl1 Product ooo atom_seqs_of fun index_for_bound_var _ [_] 0 = 0 | index_for_bound_var card (_ :: Ts) 0 = index_for_bound_var card Ts 0 + arity_of S_Rep card (hd Ts) | index_for_bound_var card Ts n = index_for_bound_var card (tl Ts) (n - 1) fun vars_for_bound_var card R Ts j = map (curry Var 1) (index_seq (index_for_bound_var card Ts j) (arity_of R card (nth Ts j))) val rel_expr_for_bound_var = foldl1 Product oooo vars_for_bound_var fun decls_for R card Ts T = map2 (curry DeclOne o pair 1) (index_seq (index_for_bound_var card (T :: Ts) 0) (arity_of R card (nth (T :: Ts) 0))) (atom_seqs_of R card T) val atom_product = foldl1 Product o map Atom val false_atom_num = 0 val true_atom_num = 1 val false_atom = Atom false_atom_num val true_atom = Atom true_atom_num fun kodkod_formula_from_term ctxt total card complete concrete frees = let fun F_from_S_rep (SOME false) r = Not (RelEq (r, false_atom)) | F_from_S_rep _ r = RelEq (r, true_atom) fun S_rep_from_F NONE f = RelIf (f, true_atom, false_atom) | S_rep_from_F (SOME true) f = RelIf (f, true_atom, None) | S_rep_from_F (SOME false) f = RelIf (Not f, false_atom, None) fun R_rep_from_S_rep (Type (\<^type_name>\fun\, [T1, T2])) r = if total andalso T2 = bool_T then let val jss = atom_schema_of S_Rep card T1 |> map (rpair 0) |> all_combinations in map2 (fn i => fn js => (* RelIf (F_from_S_rep NONE (Project (r, [Num i])), atom_product js, empty_n_ary_rel (length js)) *) Join (Project (r, [Num i]), atom_product (false_atom_num :: js)) ) (index_seq 0 (length jss)) jss |> foldl1 Union end else let val jss = atom_schema_of S_Rep card T1 |> map (rpair 0) |> all_combinations val arity2 = arity_of S_Rep card T2 in map2 (fn i => fn js => Product (atom_product js, Project (r, num_seq (i * arity2) arity2) |> R_rep_from_S_rep T2)) (index_seq 0 (length jss)) jss |> foldl1 Union end | R_rep_from_S_rep _ r = r fun S_rep_from_R_rep Ts (T as Type (\<^type_name>\fun\, _)) r = Comprehension (decls_for S_Rep card Ts T, RelEq (R_rep_from_S_rep T (rel_expr_for_bound_var card S_Rep (T :: Ts) 0), r)) | S_rep_from_R_rep _ _ r = r fun partial_eq pos Ts (Type (\<^type_name>\fun\, [T1, T2])) t1 t2 = HOLogic.mk_all ("x", T1, HOLogic.eq_const T2 $ (incr_boundvars 1 t1 $ Bound 0) $ (incr_boundvars 1 t2 $ Bound 0)) |> to_F (SOME pos) Ts | partial_eq pos Ts T t1 t2 = if pos andalso not (concrete T) then False else (t1, t2) |> apply2 (to_R_rep Ts) |> (if pos then Some o Intersect else Lone o Union) and to_F pos Ts t = (case t of \<^const>\Not\ $ t1 => Not (to_F (Option.map not pos) Ts t1) | \<^const>\False\ => False | \<^const>\True\ => True | Const (\<^const_name>\All\, _) $ Abs (_, T, t') => if pos = SOME true andalso not (complete T) then False else All (decls_for S_Rep card Ts T, to_F pos (T :: Ts) t') | (t0 as Const (\<^const_name>\All\, _)) $ t1 => to_F pos Ts (t0 $ eta_expand Ts t1 1) | Const (\<^const_name>\Ex\, _) $ Abs (_, T, t') => if pos = SOME false andalso not (complete T) then True else Exist (decls_for S_Rep card Ts T, to_F pos (T :: Ts) t') | (t0 as Const (\<^const_name>\Ex\, _)) $ t1 => to_F pos Ts (t0 $ eta_expand Ts t1 1) | Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ t1 $ t2 => (case pos of NONE => RelEq (to_R_rep Ts t1, to_R_rep Ts t2) | SOME pos => partial_eq pos Ts T t1 t2) | Const (\<^const_name>\ord_class.less_eq\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [T', \<^typ>\bool\]), _])) $ t1 $ t2 => (case pos of NONE => Subset (to_R_rep Ts t1, to_R_rep Ts t2) | SOME true => Subset (Difference (atom_seq_product_of S_Rep card T', Join (to_R_rep Ts t1, false_atom)), Join (to_R_rep Ts t2, true_atom)) | SOME false => Subset (Join (to_R_rep Ts t1, true_atom), Difference (atom_seq_product_of S_Rep card T', Join (to_R_rep Ts t2, false_atom)))) | \<^const>\HOL.conj\ $ t1 $ t2 => And (to_F pos Ts t1, to_F pos Ts t2) | \<^const>\HOL.disj\ $ t1 $ t2 => Or (to_F pos Ts t1, to_F pos Ts t2) | \<^const>\HOL.implies\ $ t1 $ t2 => Implies (to_F (Option.map not pos) Ts t1, to_F pos Ts t2) | Const (\<^const_name>\Set.member\, _) $ t1 $ t2 => to_F pos Ts (t2 $ t1) | t1 $ t2 => (case pos of NONE => Subset (to_S_rep Ts t2, to_R_rep Ts t1) | SOME pos => let val kt1 = to_R_rep Ts t1 val kt2 = to_S_rep Ts t2 val kT = atom_seq_product_of S_Rep card (fastype_of1 (Ts, t2)) in if pos then Not (Subset (kt2, Difference (kT, Join (kt1, true_atom)))) else Subset (kt2, Difference (kT, Join (kt1, false_atom))) end) | _ => raise SAME ()) handle SAME () => F_from_S_rep pos (to_R_rep Ts t) and to_S_rep Ts t = case t of Const (\<^const_name>\Pair\, _) $ t1 $ t2 => Product (to_S_rep Ts t1, to_S_rep Ts t2) | Const (\<^const_name>\Pair\, _) $ _ => to_S_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\Pair\, _) => to_S_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\fst\, _) $ t1 => let val fst_arity = arity_of S_Rep card (fastype_of1 (Ts, t)) in Project (to_S_rep Ts t1, num_seq 0 fst_arity) end | Const (\<^const_name>\fst\, _) => to_S_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\snd\, _) $ t1 => let val pair_arity = arity_of S_Rep card (fastype_of1 (Ts, t1)) val snd_arity = arity_of S_Rep card (fastype_of1 (Ts, t)) val fst_arity = pair_arity - snd_arity in Project (to_S_rep Ts t1, num_seq fst_arity snd_arity) end | Const (\<^const_name>\snd\, _) => to_S_rep Ts (eta_expand Ts t 1) | Bound j => rel_expr_for_bound_var card S_Rep Ts j | _ => S_rep_from_R_rep Ts (fastype_of1 (Ts, t)) (to_R_rep Ts t) and partial_set_op swap1 swap2 op1 op2 Ts t1 t2 = let val kt1 = to_R_rep Ts t1 val kt2 = to_R_rep Ts t2 val (a11, a21) = (false_atom, true_atom) |> swap1 ? swap val (a12, a22) = (false_atom, true_atom) |> swap2 ? swap in Union (Product (op1 (Join (kt1, a11), Join (kt2, a12)), true_atom), Product (op2 (Join (kt1, a21), Join (kt2, a22)), false_atom)) end and to_R_rep Ts t = (case t of \<^const>\Not\ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\All\, _) => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\Ex\, _) => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\HOL.eq\, _) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\HOL.eq\, _) => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\ord_class.less_eq\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [_, \<^typ>\bool\]), _])) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\ord_class.less_eq\, _) => to_R_rep Ts (eta_expand Ts t 2) | \<^const>\HOL.conj\ $ _ => to_R_rep Ts (eta_expand Ts t 1) | \<^const>\HOL.conj\ => to_R_rep Ts (eta_expand Ts t 2) | \<^const>\HOL.disj\ $ _ => to_R_rep Ts (eta_expand Ts t 1) | \<^const>\HOL.disj\ => to_R_rep Ts (eta_expand Ts t 2) | \<^const>\HOL.implies\ $ _ => to_R_rep Ts (eta_expand Ts t 1) | \<^const>\HOL.implies\ => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\Set.member\, _) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\Set.member\, _) => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\Collect\, _) $ t' => to_R_rep Ts t' | Const (\<^const_name>\Collect\, _) => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\bot_class.bot\, T as Type (\<^type_name>\fun\, [T', \<^typ>\bool\])) => if total then empty_n_ary_rel (arity_of (R_Rep total) card T) else Product (atom_seq_product_of (R_Rep total) card T', false_atom) | Const (\<^const_name>\top_class.top\, T as Type (\<^type_name>\fun\, [T', \<^typ>\bool\])) => if total then atom_seq_product_of (R_Rep total) card T else Product (atom_seq_product_of (R_Rep total) card T', true_atom) | Const (\<^const_name>\insert\, Type (_, [T, _])) $ t1 $ t2 => if total then Union (to_S_rep Ts t1, to_R_rep Ts t2) else let val kt1 = to_S_rep Ts t1 val kt2 = to_R_rep Ts t2 in RelIf (Some kt1, if arity_of S_Rep card T = 1 then Override (kt2, Product (kt1, true_atom)) else Union (Difference (kt2, Product (kt1, false_atom)), Product (kt1, true_atom)), Difference (kt2, Product (atom_seq_product_of S_Rep card T, false_atom))) end | Const (\<^const_name>\insert\, _) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\insert\, _) => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\trancl\, Type (_, [Type (_, [Type (_, [T', _]), _]), _])) $ t1 => if arity_of S_Rep card T' = 1 then if total then Closure (to_R_rep Ts t1) else let val kt1 = to_R_rep Ts t1 val true_core_kt = Closure (Join (kt1, true_atom)) val kTx = atom_seq_product_of S_Rep card (HOLogic.mk_prodT (`I T')) val false_mantle_kt = Difference (kTx, Closure (Difference (kTx, Join (kt1, false_atom)))) in Union (Product (Difference (false_mantle_kt, true_core_kt), false_atom), Product (true_core_kt, true_atom)) end else error "Not supported: Transitive closure for function or pair type." | Const (\<^const_name>\trancl\, _) => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\inf_class.inf\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [_, \<^typ>\bool\]), _])) $ t1 $ t2 => if total then Intersect (to_R_rep Ts t1, to_R_rep Ts t2) else partial_set_op true true Intersect Union Ts t1 t2 | Const (\<^const_name>\inf_class.inf\, _) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\inf_class.inf\, _) => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\sup_class.sup\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [_, \<^typ>\bool\]), _])) $ t1 $ t2 => if total then Union (to_R_rep Ts t1, to_R_rep Ts t2) else partial_set_op true true Union Intersect Ts t1 t2 | Const (\<^const_name>\sup_class.sup\, _) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\sup_class.sup\, _) => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\minus_class.minus\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [_, \<^typ>\bool\]), _])) $ t1 $ t2 => if total then Difference (to_R_rep Ts t1, to_R_rep Ts t2) else partial_set_op true false Intersect Union Ts t1 t2 | Const (\<^const_name>\minus_class.minus\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [_, \<^typ>\bool\]), _])) $ _ => to_R_rep Ts (eta_expand Ts t 1) | Const (\<^const_name>\minus_class.minus\, Type (\<^type_name>\fun\, [Type (\<^type_name>\fun\, [_, \<^typ>\bool\]), _])) => to_R_rep Ts (eta_expand Ts t 2) | Const (\<^const_name>\Pair\, _) $ _ $ _ => to_S_rep Ts t | Const (\<^const_name>\Pair\, _) $ _ => to_S_rep Ts t | Const (\<^const_name>\Pair\, _) => to_S_rep Ts t | Const (\<^const_name>\fst\, _) $ _ => raise SAME () | Const (\<^const_name>\fst\, _) => raise SAME () | Const (\<^const_name>\snd\, _) $ _ => raise SAME () | Const (\<^const_name>\snd\, _) => raise SAME () | \<^const>\False\ => false_atom | \<^const>\True\ => true_atom | Free (x as (_, T)) => Rel (arity_of (R_Rep total) card T, find_index (curry (op =) x) frees) | Term.Var _ => error "Not supported: Schematic variables." | Bound _ => raise SAME () | Abs (_, T, t') => (case (total, fastype_of1 (T :: Ts, t')) of (true, \<^typ>\bool\) => Comprehension (decls_for S_Rep card Ts T, to_F NONE (T :: Ts) t') | (_, T') => Comprehension (decls_for S_Rep card Ts T @ decls_for (R_Rep total) card (T :: Ts) T', Subset (rel_expr_for_bound_var card (R_Rep total) (T' :: T :: Ts) 0, to_R_rep (T :: Ts) t'))) | t1 $ t2 => (case fastype_of1 (Ts, t) of \<^typ>\bool\ => if total then S_rep_from_F NONE (to_F NONE Ts t) else RelIf (to_F (SOME true) Ts t, true_atom, RelIf (Not (to_F (SOME false) Ts t), false_atom, None)) | T => let val T2 = fastype_of1 (Ts, t2) in case arity_of S_Rep card T2 of 1 => Join (to_S_rep Ts t2, to_R_rep Ts t1) | arity2 => let val res_arity = arity_of (R_Rep total) card T in Project (Intersect (Product (to_S_rep Ts t2, atom_seq_product_of (R_Rep total) card T), to_R_rep Ts t1), num_seq arity2 res_arity) end end) | _ => error ("Not supported: Term " ^ quote (Syntax.string_of_term ctxt t) ^ ".")) handle SAME () => R_rep_from_S_rep (fastype_of1 (Ts, t)) (to_S_rep Ts t) in to_F (if total then NONE else SOME true) [] end fun bound_for_free total card i (s, T) = let val js = atom_schema_of (R_Rep total) card T in ([((length js, i), s)], [TupleSet [], atom_schema_of (R_Rep total) card T |> map (rpair 0) |> tuple_set_from_atom_schema]) end fun declarative_axiom_for_rel_expr total card Ts (Type (\<^type_name>\fun\, [T1, T2])) r = if total andalso body_type T2 = bool_T then True else All (decls_for S_Rep card Ts T1, declarative_axiom_for_rel_expr total card (T1 :: Ts) T2 (List.foldl Join r (vars_for_bound_var card S_Rep (T1 :: Ts) 0))) | declarative_axiom_for_rel_expr total _ _ _ r = (if total then One else Lone) r fun declarative_axiom_for_free total card i (_, T) = declarative_axiom_for_rel_expr total card [] T (Rel (arity_of (R_Rep total) card T, i)) (* Hack to make the old code work as is with sets. *) fun unsetify_type (Type (\<^type_name>\set\, [T])) = unsetify_type T --> bool_T | unsetify_type (Type (s, Ts)) = Type (s, map unsetify_type Ts) | unsetify_type T = T fun kodkod_problem_from_term ctxt total raw_card raw_infinite t = let val thy = Proof_Context.theory_of ctxt fun card (Type (\<^type_name>\fun\, [T1, T2])) = reasonable_power (card T2) (card T1) | card (Type (\<^type_name>\prod\, [T1, T2])) = card T1 * card T2 | card \<^typ>\bool\ = 2 | card T = Int.max (1, raw_card T) fun complete (Type (\<^type_name>\fun\, [T1, T2])) = concrete T1 andalso complete T2 | complete (Type (\<^type_name>\prod\, Ts)) = forall complete Ts | complete T = not (raw_infinite T) and concrete (Type (\<^type_name>\fun\, [T1, T2])) = complete T1 andalso concrete T2 | concrete (Type (\<^type_name>\prod\, Ts)) = forall concrete Ts | concrete _ = true val neg_t = \<^const>\Not\ $ Object_Logic.atomize_term ctxt t |> map_types unsetify_type val _ = fold_types (K o check_type ctxt raw_infinite) neg_t () val frees = Term.add_frees neg_t [] val bounds = map2 (bound_for_free total card) (index_seq 0 (length frees)) frees val declarative_axioms = map2 (declarative_axiom_for_free total card) (index_seq 0 (length frees)) frees val formula = neg_t |> kodkod_formula_from_term ctxt total card complete concrete frees |> fold_rev (curry And) declarative_axioms val univ_card = univ_card 0 0 0 bounds formula in {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [], bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula} end fun solve_any_kodkod_problem thy problems = let val {debug, overlord, timeout, ...} = Nitpick_Commands.default_params thy [] val deadline = Time.now () + timeout val max_threads = 1 val max_solutions = 1 in case solve_any_problem debug overlord deadline max_threads max_solutions problems of - JavaNotFound => "unknown" - | Normal ([], _, _) => "none" + Normal ([], _, _) => "none" | Normal _ => "genuine" | TimedOut _ => "unknown" | Error (s, _) => error ("Kodkod error: " ^ s) end val default_raw_infinite = member (op =) [\<^typ>\nat\, \<^typ>\int\] fun minipick ctxt n t = let val thy = Proof_Context.theory_of ctxt val {total_consts, ...} = Nitpick_Commands.default_params thy [] val totals = total_consts |> Option.map single |> the_default [true, false] fun problem_for (total, k) = kodkod_problem_from_term ctxt total (K k) default_raw_infinite t in (totals, 1 upto n) |-> map_product pair |> map problem_for |> solve_any_kodkod_problem (Proof_Context.theory_of ctxt) end fun minipick_expect ctxt expect n t = if getenv "KODKODI" <> "" then if minipick ctxt n t = expect then () else error ("\"minipick_expect\" expected " ^ quote expect) else () end; diff --git a/src/HOL/Tools/Nitpick/kodkod.ML b/src/HOL/Tools/Nitpick/kodkod.ML --- a/src/HOL/Tools/Nitpick/kodkod.ML +++ b/src/HOL/Tools/Nitpick/kodkod.ML @@ -1,1108 +1,1104 @@ (* Title: HOL/Tools/Nitpick/kodkod.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2008-2014 ML interface for Kodkod. *) signature KODKOD = sig type n_ary_index = int * int type setting = string * string datatype tuple = Tuple of int list | TupleIndex of n_ary_index | TupleReg of n_ary_index datatype tuple_set = TupleUnion of tuple_set * tuple_set | TupleDifference of tuple_set * tuple_set | TupleIntersect of tuple_set * tuple_set | TupleProduct of tuple_set * tuple_set | TupleProject of tuple_set * int | TupleSet of tuple list | TupleRange of tuple * tuple | TupleArea of tuple * tuple | TupleAtomSeq of int * int | TupleSetReg of n_ary_index datatype tuple_assign = AssignTuple of n_ary_index * tuple | AssignTupleSet of n_ary_index * tuple_set type bound = (n_ary_index * string) list * tuple_set list type int_bound = int option * tuple_set list datatype formula = All of decl list * formula | Exist of decl list * formula | FormulaLet of expr_assign list * formula | FormulaIf of formula * formula * formula | Or of formula * formula | Iff of formula * formula | Implies of formula * formula | And of formula * formula | Not of formula | Acyclic of n_ary_index | Function of n_ary_index * rel_expr * rel_expr | Functional of n_ary_index * rel_expr * rel_expr | TotalOrdering of n_ary_index * rel_expr * rel_expr * rel_expr | Subset of rel_expr * rel_expr | RelEq of rel_expr * rel_expr | IntEq of int_expr * int_expr | LT of int_expr * int_expr | LE of int_expr * int_expr | No of rel_expr | Lone of rel_expr | One of rel_expr | Some of rel_expr | False | True | FormulaReg of int and rel_expr = RelLet of expr_assign list * rel_expr | RelIf of formula * rel_expr * rel_expr | Union of rel_expr * rel_expr | Difference of rel_expr * rel_expr | Override of rel_expr * rel_expr | Intersect of rel_expr * rel_expr | Product of rel_expr * rel_expr | IfNo of rel_expr * rel_expr | Project of rel_expr * int_expr list | Join of rel_expr * rel_expr | Closure of rel_expr | ReflexiveClosure of rel_expr | Transpose of rel_expr | Comprehension of decl list * formula | Bits of int_expr | Int of int_expr | Iden | Ints | None | Univ | Atom of int | AtomSeq of int * int | Rel of n_ary_index | Var of n_ary_index | RelReg of n_ary_index and int_expr = Sum of decl list * int_expr | IntLet of expr_assign list * int_expr | IntIf of formula * int_expr * int_expr | SHL of int_expr * int_expr | SHA of int_expr * int_expr | SHR of int_expr * int_expr | Add of int_expr * int_expr | Sub of int_expr * int_expr | Mult of int_expr * int_expr | Div of int_expr * int_expr | Mod of int_expr * int_expr | Cardinality of rel_expr | SetSum of rel_expr | BitOr of int_expr * int_expr | BitXor of int_expr * int_expr | BitAnd of int_expr * int_expr | BitNot of int_expr | Neg of int_expr | Absolute of int_expr | Signum of int_expr | Num of int | IntReg of int and decl = DeclNo of n_ary_index * rel_expr | DeclLone of n_ary_index * rel_expr | DeclOne of n_ary_index * rel_expr | DeclSome of n_ary_index * rel_expr | DeclSet of n_ary_index * rel_expr and expr_assign = AssignFormulaReg of int * formula | AssignRelReg of n_ary_index * rel_expr | AssignIntReg of int * int_expr type 'a fold_expr_funcs = {formula_func: formula -> 'a -> 'a, rel_expr_func: rel_expr -> 'a -> 'a, int_expr_func: int_expr -> 'a -> 'a} val kodkodi_version : unit -> int list val fold_formula : 'a fold_expr_funcs -> formula -> 'a -> 'a val fold_rel_expr : 'a fold_expr_funcs -> rel_expr -> 'a -> 'a val fold_int_expr : 'a fold_expr_funcs -> int_expr -> 'a -> 'a val fold_decl : 'a fold_expr_funcs -> decl -> 'a -> 'a val fold_expr_assign : 'a fold_expr_funcs -> expr_assign -> 'a -> 'a type 'a fold_tuple_funcs = {tuple_func: tuple -> 'a -> 'a, tuple_set_func: tuple_set -> 'a -> 'a} val fold_tuple : 'a fold_tuple_funcs -> tuple -> 'a -> 'a val fold_tuple_set : 'a fold_tuple_funcs -> tuple_set -> 'a -> 'a val fold_tuple_assign : 'a fold_tuple_funcs -> tuple_assign -> 'a -> 'a val fold_bound : 'a fold_expr_funcs -> 'a fold_tuple_funcs -> bound -> 'a -> 'a val fold_int_bound : 'a fold_tuple_funcs -> int_bound -> 'a -> 'a type problem = {comment: string, settings: setting list, univ_card: int, tuple_assigns: tuple_assign list, bounds: bound list, int_bounds: int_bound list, expr_assigns: expr_assign list, formula: formula} type raw_bound = n_ary_index * int list list datatype outcome = - JavaNotFound | Normal of (int * raw_bound list) list * int list * string | TimedOut of int list | Error of string * int list exception SYNTAX of string * string val max_arity : int -> int val arity_of_rel_expr : rel_expr -> int val is_problem_trivially_false : problem -> bool val problems_equivalent : problem * problem -> bool val solve_any_problem : bool -> bool -> Time.time -> int -> int -> problem list -> outcome end; structure Kodkod : KODKOD = struct type n_ary_index = int * int type setting = string * string datatype tuple = Tuple of int list | TupleIndex of n_ary_index | TupleReg of n_ary_index datatype tuple_set = TupleUnion of tuple_set * tuple_set | TupleDifference of tuple_set * tuple_set | TupleIntersect of tuple_set * tuple_set | TupleProduct of tuple_set * tuple_set | TupleProject of tuple_set * int | TupleSet of tuple list | TupleRange of tuple * tuple | TupleArea of tuple * tuple | TupleAtomSeq of int * int | TupleSetReg of n_ary_index datatype tuple_assign = AssignTuple of n_ary_index * tuple | AssignTupleSet of n_ary_index * tuple_set type bound = (n_ary_index * string) list * tuple_set list type int_bound = int option * tuple_set list datatype formula = All of decl list * formula | Exist of decl list * formula | FormulaLet of expr_assign list * formula | FormulaIf of formula * formula * formula | Or of formula * formula | Iff of formula * formula | Implies of formula * formula | And of formula * formula | Not of formula | Acyclic of n_ary_index | Function of n_ary_index * rel_expr * rel_expr | Functional of n_ary_index * rel_expr * rel_expr | TotalOrdering of n_ary_index * rel_expr * rel_expr * rel_expr | Subset of rel_expr * rel_expr | RelEq of rel_expr * rel_expr | IntEq of int_expr * int_expr | LT of int_expr * int_expr | LE of int_expr * int_expr | No of rel_expr | Lone of rel_expr | One of rel_expr | Some of rel_expr | False | True | FormulaReg of int and rel_expr = RelLet of expr_assign list * rel_expr | RelIf of formula * rel_expr * rel_expr | Union of rel_expr * rel_expr | Difference of rel_expr * rel_expr | Override of rel_expr * rel_expr | Intersect of rel_expr * rel_expr | Product of rel_expr * rel_expr | IfNo of rel_expr * rel_expr | Project of rel_expr * int_expr list | Join of rel_expr * rel_expr | Closure of rel_expr | ReflexiveClosure of rel_expr | Transpose of rel_expr | Comprehension of decl list * formula | Bits of int_expr | Int of int_expr | Iden | Ints | None | Univ | Atom of int | AtomSeq of int * int | Rel of n_ary_index | Var of n_ary_index | RelReg of n_ary_index and int_expr = Sum of decl list * int_expr | IntLet of expr_assign list * int_expr | IntIf of formula * int_expr * int_expr | SHL of int_expr * int_expr | SHA of int_expr * int_expr | SHR of int_expr * int_expr | Add of int_expr * int_expr | Sub of int_expr * int_expr | Mult of int_expr * int_expr | Div of int_expr * int_expr | Mod of int_expr * int_expr | Cardinality of rel_expr | SetSum of rel_expr | BitOr of int_expr * int_expr | BitXor of int_expr * int_expr | BitAnd of int_expr * int_expr | BitNot of int_expr | Neg of int_expr | Absolute of int_expr | Signum of int_expr | Num of int | IntReg of int and decl = DeclNo of n_ary_index * rel_expr | DeclLone of n_ary_index * rel_expr | DeclOne of n_ary_index * rel_expr | DeclSome of n_ary_index * rel_expr | DeclSet of n_ary_index * rel_expr and expr_assign = AssignFormulaReg of int * formula | AssignRelReg of n_ary_index * rel_expr | AssignIntReg of int * int_expr type problem = {comment: string, settings: setting list, univ_card: int, tuple_assigns: tuple_assign list, bounds: bound list, int_bounds: int_bound list, expr_assigns: expr_assign list, formula: formula} type raw_bound = n_ary_index * int list list datatype outcome = - JavaNotFound | Normal of (int * raw_bound list) list * int list * string | TimedOut of int list | Error of string * int list exception SYNTAX of string * string type 'a fold_expr_funcs = {formula_func: formula -> 'a -> 'a, rel_expr_func: rel_expr -> 'a -> 'a, int_expr_func: int_expr -> 'a -> 'a} fun kodkodi_version () = getenv "KODKODI_VERSION" |> space_explode "." |> map (the_default 0 o Int.fromString) (** Auxiliary functions on Kodkod problems **) fun fold_formula (F : 'a fold_expr_funcs) formula = case formula of All (ds, f) => fold (fold_decl F) ds #> fold_formula F f | Exist (ds, f) => fold (fold_decl F) ds #> fold_formula F f | FormulaLet (bs, f) => fold (fold_expr_assign F) bs #> fold_formula F f | FormulaIf (f, f1, f2) => fold_formula F f #> fold_formula F f1 #> fold_formula F f2 | Or (f1, f2) => fold_formula F f1 #> fold_formula F f2 | Iff (f1, f2) => fold_formula F f1 #> fold_formula F f2 | Implies (f1, f2) => fold_formula F f1 #> fold_formula F f2 | And (f1, f2) => fold_formula F f1 #> fold_formula F f2 | Not f => fold_formula F f | Acyclic x => fold_rel_expr F (Rel x) | Function (x, r1, r2) => fold_rel_expr F (Rel x) #> fold_rel_expr F r1 #> fold_rel_expr F r2 | Functional (x, r1, r2) => fold_rel_expr F (Rel x) #> fold_rel_expr F r1 #> fold_rel_expr F r2 | TotalOrdering (x, r1, r2, r3) => fold_rel_expr F (Rel x) #> fold_rel_expr F r1 #> fold_rel_expr F r2 #> fold_rel_expr F r3 | Subset (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | RelEq (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | IntEq (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | LT (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | LE (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | No r => fold_rel_expr F r | Lone r => fold_rel_expr F r | One r => fold_rel_expr F r | Some r => fold_rel_expr F r | False => #formula_func F formula | True => #formula_func F formula | FormulaReg _ => #formula_func F formula and fold_rel_expr F rel_expr = case rel_expr of RelLet (bs, r) => fold (fold_expr_assign F) bs #> fold_rel_expr F r | RelIf (f, r1, r2) => fold_formula F f #> fold_rel_expr F r1 #> fold_rel_expr F r2 | Union (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | Difference (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | Override (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | Intersect (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | Product (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | IfNo (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | Project (r1, is) => fold_rel_expr F r1 #> fold (fold_int_expr F) is | Join (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 | Closure r => fold_rel_expr F r | ReflexiveClosure r => fold_rel_expr F r | Transpose r => fold_rel_expr F r | Comprehension (ds, f) => fold (fold_decl F) ds #> fold_formula F f | Bits i => fold_int_expr F i | Int i => fold_int_expr F i | Iden => #rel_expr_func F rel_expr | Ints => #rel_expr_func F rel_expr | None => #rel_expr_func F rel_expr | Univ => #rel_expr_func F rel_expr | Atom _ => #rel_expr_func F rel_expr | AtomSeq _ => #rel_expr_func F rel_expr | Rel _ => #rel_expr_func F rel_expr | Var _ => #rel_expr_func F rel_expr | RelReg _ => #rel_expr_func F rel_expr and fold_int_expr F int_expr = case int_expr of Sum (ds, i) => fold (fold_decl F) ds #> fold_int_expr F i | IntLet (bs, i) => fold (fold_expr_assign F) bs #> fold_int_expr F i | IntIf (f, i1, i2) => fold_formula F f #> fold_int_expr F i1 #> fold_int_expr F i2 | SHL (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | SHA (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | SHR (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | Add (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | Sub (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | Mult (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | Div (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | Mod (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | Cardinality r => fold_rel_expr F r | SetSum r => fold_rel_expr F r | BitOr (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | BitXor (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | BitAnd (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 | BitNot i => fold_int_expr F i | Neg i => fold_int_expr F i | Absolute i => fold_int_expr F i | Signum i => fold_int_expr F i | Num _ => #int_expr_func F int_expr | IntReg _ => #int_expr_func F int_expr and fold_decl F decl = case decl of DeclNo (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r | DeclLone (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r | DeclOne (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r | DeclSome (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r | DeclSet (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r and fold_expr_assign F assign = case assign of AssignFormulaReg (x, f) => fold_formula F (FormulaReg x) #> fold_formula F f | AssignRelReg (x, r) => fold_rel_expr F (RelReg x) #> fold_rel_expr F r | AssignIntReg (x, i) => fold_int_expr F (IntReg x) #> fold_int_expr F i type 'a fold_tuple_funcs = {tuple_func: tuple -> 'a -> 'a, tuple_set_func: tuple_set -> 'a -> 'a} fun fold_tuple (F : 'a fold_tuple_funcs) = #tuple_func F fun fold_tuple_set F tuple_set = case tuple_set of TupleUnion (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 | TupleDifference (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 | TupleIntersect (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 | TupleProduct (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 | TupleProject (ts, _) => fold_tuple_set F ts | TupleSet ts => fold (fold_tuple F) ts | TupleRange (t1, t2) => fold_tuple F t1 #> fold_tuple F t2 | TupleArea (t1, t2) => fold_tuple F t1 #> fold_tuple F t2 | TupleAtomSeq _ => #tuple_set_func F tuple_set | TupleSetReg _ => #tuple_set_func F tuple_set fun fold_tuple_assign F assign = case assign of AssignTuple (x, t) => fold_tuple F (TupleReg x) #> fold_tuple F t | AssignTupleSet (x, ts) => fold_tuple_set F (TupleSetReg x) #> fold_tuple_set F ts fun fold_bound expr_F tuple_F (zs, tss) = fold (fold_rel_expr expr_F) (map (Rel o fst) zs) #> fold (fold_tuple_set tuple_F) tss fun fold_int_bound F (_, tss) = fold (fold_tuple_set F) tss fun max_arity univ_card = floor (Math.ln 2147483647.0 / Math.ln (Real.fromInt univ_card)) fun arity_of_rel_expr (RelLet (_, r)) = arity_of_rel_expr r | arity_of_rel_expr (RelIf (_, r1, _)) = arity_of_rel_expr r1 | arity_of_rel_expr (Union (r1, _)) = arity_of_rel_expr r1 | arity_of_rel_expr (Difference (r1, _)) = arity_of_rel_expr r1 | arity_of_rel_expr (Override (r1, _)) = arity_of_rel_expr r1 | arity_of_rel_expr (Intersect (r1, _)) = arity_of_rel_expr r1 | arity_of_rel_expr (Product (r1, r2)) = sum_arities_of_rel_exprs r1 r2 | arity_of_rel_expr (IfNo (r1, _)) = arity_of_rel_expr r1 | arity_of_rel_expr (Project (_, is)) = length is | arity_of_rel_expr (Join (r1, r2)) = sum_arities_of_rel_exprs r1 r2 - 2 | arity_of_rel_expr (Closure _) = 2 | arity_of_rel_expr (ReflexiveClosure _) = 2 | arity_of_rel_expr (Transpose _) = 2 | arity_of_rel_expr (Comprehension (ds, _)) = fold (curry op + o arity_of_decl) ds 0 | arity_of_rel_expr (Bits _) = 1 | arity_of_rel_expr (Int _) = 1 | arity_of_rel_expr Iden = 2 | arity_of_rel_expr Ints = 1 | arity_of_rel_expr None = 1 | arity_of_rel_expr Univ = 1 | arity_of_rel_expr (Atom _) = 1 | arity_of_rel_expr (AtomSeq _) = 1 | arity_of_rel_expr (Rel (n, _)) = n | arity_of_rel_expr (Var (n, _)) = n | arity_of_rel_expr (RelReg (n, _)) = n and sum_arities_of_rel_exprs r1 r2 = arity_of_rel_expr r1 + arity_of_rel_expr r2 and arity_of_decl (DeclNo ((n, _), _)) = n | arity_of_decl (DeclLone ((n, _), _)) = n | arity_of_decl (DeclOne ((n, _), _)) = n | arity_of_decl (DeclSome ((n, _), _)) = n | arity_of_decl (DeclSet ((n, _), _)) = n fun is_problem_trivially_false ({formula = False, ...} : problem) = true | is_problem_trivially_false _ = false val chop_solver = take 2 o space_explode "," fun settings_equivalent ([], []) = true | settings_equivalent ((key1, value1) :: settings1, (key2, value2) :: settings2) = key1 = key2 andalso (value1 = value2 orelse key1 = "delay" orelse (key1 = "solver" andalso chop_solver value1 = chop_solver value2)) andalso settings_equivalent (settings1, settings2) | settings_equivalent _ = false fun problems_equivalent (p1 : problem, p2 : problem) = #univ_card p1 = #univ_card p2 andalso #formula p1 = #formula p2 andalso #bounds p1 = #bounds p2 andalso #expr_assigns p1 = #expr_assigns p2 andalso #tuple_assigns p1 = #tuple_assigns p2 andalso #int_bounds p1 = #int_bounds p2 andalso settings_equivalent (#settings p1, #settings p2) (** Serialization of problem **) fun base_name j = if j < 0 then string_of_int (~j - 1) ^ "'" else string_of_int j fun n_ary_name (1, j) prefix _ _ = prefix ^ base_name j | n_ary_name (2, j) _ prefix _ = prefix ^ base_name j | n_ary_name (n, j) _ _ prefix = prefix ^ string_of_int n ^ "_" ^ base_name j fun atom_name j = "A" ^ base_name j fun atom_seq_name (k, 0) = "u" ^ base_name k | atom_seq_name (k, j0) = "u" ^ base_name k ^ "@" ^ base_name j0 fun formula_reg_name j = "$f" ^ base_name j fun rel_reg_name j = "$e" ^ base_name j fun int_reg_name j = "$i" ^ base_name j fun tuple_name x = n_ary_name x "A" "P" "T" fun rel_name x = n_ary_name x "s" "r" "m" fun var_name x = n_ary_name x "S" "R" "M" fun tuple_reg_name x = n_ary_name x "$A" "$P" "$T" fun tuple_set_reg_name x = n_ary_name x "$a" "$p" "$t" fun inline_comment "" = "" | inline_comment comment = " /* " ^ translate_string (fn "\n" => " " | "*" => "* " | s => s) comment ^ " */" fun block_comment "" = "" | block_comment comment = prefix_lines "// " comment ^ "\n" fun commented_rel_name (x, s) = rel_name x ^ inline_comment s fun string_for_tuple (Tuple js) = "[" ^ commas (map atom_name js) ^ "]" | string_for_tuple (TupleIndex x) = tuple_name x | string_for_tuple (TupleReg x) = tuple_reg_name x val no_prec = 100 val prec_TupleUnion = 1 val prec_TupleIntersect = 2 val prec_TupleProduct = 3 val prec_TupleProject = 4 fun precedence_ts (TupleUnion _) = prec_TupleUnion | precedence_ts (TupleDifference _) = prec_TupleUnion | precedence_ts (TupleIntersect _) = prec_TupleIntersect | precedence_ts (TupleProduct _) = prec_TupleProduct | precedence_ts (TupleProject _) = prec_TupleProject | precedence_ts _ = no_prec fun string_for_tuple_set tuple_set = let fun sub tuple_set outer_prec = let val prec = precedence_ts tuple_set val need_parens = (prec < outer_prec) in (if need_parens then "(" else "") ^ (case tuple_set of TupleUnion (ts1, ts2) => sub ts1 prec ^ " + " ^ sub ts2 (prec + 1) | TupleDifference (ts1, ts2) => sub ts1 prec ^ " - " ^ sub ts2 (prec + 1) | TupleIntersect (ts1, ts2) => sub ts1 prec ^ " & " ^ sub ts2 prec | TupleProduct (ts1, ts2) => sub ts1 prec ^ "->" ^ sub ts2 prec | TupleProject (ts, c) => sub ts prec ^ "[" ^ string_of_int c ^ "]" | TupleSet ts => "{" ^ commas (map string_for_tuple ts) ^ "}" | TupleRange (t1, t2) => "{" ^ string_for_tuple t1 ^ (if t1 = t2 then "" else " .. " ^ string_for_tuple t2) ^ "}" | TupleArea (t1, t2) => "{" ^ string_for_tuple t1 ^ " # " ^ string_for_tuple t2 ^ "}" | TupleAtomSeq x => atom_seq_name x | TupleSetReg x => tuple_set_reg_name x) ^ (if need_parens then ")" else "") end in sub tuple_set 0 end fun string_for_tuple_assign (AssignTuple (x, t)) = tuple_reg_name x ^ " := " ^ string_for_tuple t ^ "\n" | string_for_tuple_assign (AssignTupleSet (x, ts)) = tuple_set_reg_name x ^ " := " ^ string_for_tuple_set ts ^ "\n" fun string_for_bound (zs, tss) = "bounds " ^ commas (map commented_rel_name zs) ^ ": " ^ (if length tss = 1 then "" else "[") ^ commas (map string_for_tuple_set tss) ^ (if length tss = 1 then "" else "]") ^ "\n" fun int_string_for_bound (opt_n, tss) = (case opt_n of SOME n => signed_string_of_int n ^ ": " | NONE => "") ^ "[" ^ commas (map string_for_tuple_set tss) ^ "]" val prec_All = 1 val prec_Or = 2 val prec_Iff = 3 val prec_Implies = 4 val prec_And = 5 val prec_Not = 6 val prec_Eq = 7 val prec_Some = 8 val prec_SHL = 9 val prec_Add = 10 val prec_Mult = 11 val prec_Override = 12 val prec_Intersect = 13 val prec_Product = 14 val prec_IfNo = 15 val prec_Project = 17 val prec_Join = 18 val prec_BitNot = 19 fun precedence_f (All _) = prec_All | precedence_f (Exist _) = prec_All | precedence_f (FormulaLet _) = prec_All | precedence_f (FormulaIf _) = prec_All | precedence_f (Or _) = prec_Or | precedence_f (Iff _) = prec_Iff | precedence_f (Implies _) = prec_Implies | precedence_f (And _) = prec_And | precedence_f (Not _) = prec_Not | precedence_f (Acyclic _) = no_prec | precedence_f (Function _) = no_prec | precedence_f (Functional _) = no_prec | precedence_f (TotalOrdering _) = no_prec | precedence_f (Subset _) = prec_Eq | precedence_f (RelEq _) = prec_Eq | precedence_f (IntEq _) = prec_Eq | precedence_f (LT _) = prec_Eq | precedence_f (LE _) = prec_Eq | precedence_f (No _) = prec_Some | precedence_f (Lone _) = prec_Some | precedence_f (One _) = prec_Some | precedence_f (Some _) = prec_Some | precedence_f False = no_prec | precedence_f True = no_prec | precedence_f (FormulaReg _) = no_prec and precedence_r (RelLet _) = prec_All | precedence_r (RelIf _) = prec_All | precedence_r (Union _) = prec_Add | precedence_r (Difference _) = prec_Add | precedence_r (Override _) = prec_Override | precedence_r (Intersect _) = prec_Intersect | precedence_r (Product _) = prec_Product | precedence_r (IfNo _) = prec_IfNo | precedence_r (Project _) = prec_Project | precedence_r (Join _) = prec_Join | precedence_r (Closure _) = prec_BitNot | precedence_r (ReflexiveClosure _) = prec_BitNot | precedence_r (Transpose _) = prec_BitNot | precedence_r (Comprehension _) = no_prec | precedence_r (Bits _) = no_prec | precedence_r (Int _) = no_prec | precedence_r Iden = no_prec | precedence_r Ints = no_prec | precedence_r None = no_prec | precedence_r Univ = no_prec | precedence_r (Atom _) = no_prec | precedence_r (AtomSeq _) = no_prec | precedence_r (Rel _) = no_prec | precedence_r (Var _) = no_prec | precedence_r (RelReg _) = no_prec and precedence_i (Sum _) = prec_All | precedence_i (IntLet _) = prec_All | precedence_i (IntIf _) = prec_All | precedence_i (SHL _) = prec_SHL | precedence_i (SHA _) = prec_SHL | precedence_i (SHR _) = prec_SHL | precedence_i (Add _) = prec_Add | precedence_i (Sub _) = prec_Add | precedence_i (Mult _) = prec_Mult | precedence_i (Div _) = prec_Mult | precedence_i (Mod _) = prec_Mult | precedence_i (Cardinality _) = no_prec | precedence_i (SetSum _) = no_prec | precedence_i (BitOr _) = prec_Intersect | precedence_i (BitXor _) = prec_Intersect | precedence_i (BitAnd _) = prec_Intersect | precedence_i (BitNot _) = prec_BitNot | precedence_i (Neg _) = prec_BitNot | precedence_i (Absolute _) = prec_BitNot | precedence_i (Signum _) = prec_BitNot | precedence_i (Num _) = no_prec | precedence_i (IntReg _) = no_prec fun write_problem out problems = let fun out_outmost_f (And (f1, f2)) = (out_outmost_f f1; out "\n && "; out_outmost_f f2) | out_outmost_f f = out_f f prec_And and out_f formula outer_prec = let val prec = precedence_f formula val need_parens = (prec < outer_prec) in (if need_parens then out "(" else ()); (case formula of All (ds, f) => (out "all ["; out_decls ds; out "] | "; out_f f prec) | Exist (ds, f) => (out "some ["; out_decls ds; out "] | "; out_f f prec) | FormulaLet (bs, f) => (out "let ["; out_assigns bs; out "] | "; out_f f prec) | FormulaIf (f, f1, f2) => (out "if "; out_f f prec; out " then "; out_f f1 prec; out " else "; out_f f2 prec) | Or (f1, f2) => (out_f f1 prec; out " || "; out_f f2 prec) | Iff (f1, f2) => (out_f f1 prec; out " <=> "; out_f f2 prec) | Implies (f1, f2) => (out_f f1 (prec + 1); out " => "; out_f f2 prec) | And (f1, f2) => (out_f f1 prec; out " && "; out_f f2 prec) | Not f => (out "! "; out_f f prec) | Acyclic x => out ("ACYCLIC(" ^ rel_name x ^ ")") | Function (x, r1, r2) => (out ("FUNCTION(" ^ rel_name x ^ ", "); out_r r1 0; out " -> one "; out_r r2 0; out ")") | Functional (x, r1, r2) => (out ("FUNCTION(" ^ rel_name x ^ ", "); out_r r1 0; out " -> lone "; out_r r2 0; out ")") | TotalOrdering (x, r1, r2, r3) => (out ("TOTAL_ORDERING(" ^ rel_name x ^ ", "); out_r r1 0; out ", "; out_r r2 0; out ", "; out_r r3 0; out ")") | Subset (r1, r2) => (out_r r1 prec; out " in "; out_r r2 prec) | RelEq (r1, r2) => (out_r r1 prec; out " = "; out_r r2 prec) | IntEq (i1, i2) => (out_i i1 prec; out " = "; out_i i2 prec) | LT (i1, i2) => (out_i i1 prec; out " < "; out_i i2 prec) | LE (i1, i2) => (out_i i1 prec; out " <= "; out_i i2 prec) | No r => (out "no "; out_r r prec) | Lone r => (out "lone "; out_r r prec) | One r => (out "one "; out_r r prec) | Some r => (out "some "; out_r r prec) | False => out "false" | True => out "true" | FormulaReg j => out (formula_reg_name j)); (if need_parens then out ")" else ()) end and out_r rel_expr outer_prec = let val prec = precedence_r rel_expr val need_parens = (prec < outer_prec) in (if need_parens then out "(" else ()); (case rel_expr of RelLet (bs, r) => (out "let ["; out_assigns bs; out "] | "; out_r r prec) | RelIf (f, r1, r2) => (out "if "; out_f f prec; out " then "; out_r r1 prec; out " else "; out_r r2 prec) | Union (r1, r2) => (out_r r1 prec; out " + "; out_r r2 (prec + 1)) | Difference (r1, r2) => (out_r r1 prec; out " - "; out_r r2 (prec + 1)) | Override (r1, r2) => (out_r r1 prec; out " ++ "; out_r r2 prec) | Intersect (r1, r2) => (out_r r1 prec; out " & "; out_r r2 prec) | Product (r1, r2) => (out_r r1 prec; out "->"; out_r r2 prec) | IfNo (r1, r2) => (out_r r1 prec; out "\\"; out_r r2 prec) | Project (r1, is) => (out_r r1 prec; out "["; out_columns is; out "]") | Join (r1, r2) => (out_r r1 prec; out "."; out_r r2 (prec + 1)) | Closure r => (out "^"; out_r r prec) | ReflexiveClosure r => (out "*"; out_r r prec) | Transpose r => (out "~"; out_r r prec) | Comprehension (ds, f) => (out "{["; out_decls ds; out "] | "; out_f f 0; out "}") | Bits i => (out "Bits["; out_i i 0; out "]") | Int i => (out "Int["; out_i i 0; out "]") | Iden => out "iden" | Ints => out "ints" | None => out "none" | Univ => out "univ" | Atom j => out (atom_name j) | AtomSeq x => out (atom_seq_name x) | Rel x => out (rel_name x) | Var x => out (var_name x) | RelReg (_, j) => out (rel_reg_name j)); (if need_parens then out ")" else ()) end and out_i int_expr outer_prec = let val prec = precedence_i int_expr val need_parens = (prec < outer_prec) in (if need_parens then out "(" else ()); (case int_expr of Sum (ds, i) => (out "sum ["; out_decls ds; out "] | "; out_i i prec) | IntLet (bs, i) => (out "let ["; out_assigns bs; out "] | "; out_i i prec) | IntIf (f, i1, i2) => (out "if "; out_f f prec; out " then "; out_i i1 prec; out " else "; out_i i2 prec) | SHL (i1, i2) => (out_i i1 prec; out " << "; out_i i2 (prec + 1)) | SHA (i1, i2) => (out_i i1 prec; out " >> "; out_i i2 (prec + 1)) | SHR (i1, i2) => (out_i i1 prec; out " >>> "; out_i i2 (prec + 1)) | Add (i1, i2) => (out_i i1 prec; out " + "; out_i i2 (prec + 1)) | Sub (i1, i2) => (out_i i1 prec; out " - "; out_i i2 (prec + 1)) | Mult (i1, i2) => (out_i i1 prec; out " * "; out_i i2 (prec + 1)) | Div (i1, i2) => (out_i i1 prec; out " / "; out_i i2 (prec + 1)) | Mod (i1, i2) => (out_i i1 prec; out " % "; out_i i2 (prec + 1)) | Cardinality r => (out "#("; out_r r 0; out ")") | SetSum r => (out "sum("; out_r r 0; out ")") | BitOr (i1, i2) => (out_i i1 prec; out " | "; out_i i2 prec) | BitXor (i1, i2) => (out_i i1 prec; out " ^ "; out_i i2 prec) | BitAnd (i1, i2) => (out_i i1 prec; out " & "; out_i i2 prec) | BitNot i => (out "~"; out_i i prec) | Neg i => (out "-"; out_i i prec) | Absolute i => (out "abs "; out_i i prec) | Signum i => (out "sgn "; out_i i prec) | Num k => out (signed_string_of_int k) | IntReg j => out (int_reg_name j)); (if need_parens then out ")" else ()) end and out_decls [] = () | out_decls [d] = out_decl d | out_decls (d :: ds) = (out_decl d; out ", "; out_decls ds) and out_decl (DeclNo (x, r)) = (out (var_name x); out " : no "; out_r r 0) | out_decl (DeclLone (x, r)) = (out (var_name x); out " : lone "; out_r r 0) | out_decl (DeclOne (x, r)) = (out (var_name x); out " : one "; out_r r 0) | out_decl (DeclSome (x, r)) = (out (var_name x); out " : some "; out_r r 0) | out_decl (DeclSet (x, r)) = (out (var_name x); out " : set "; out_r r 0) and out_assigns [] = () | out_assigns [b] = out_assign b | out_assigns (b :: bs) = (out_assign b; out ", "; out_assigns bs) and out_assign (AssignFormulaReg (j, f)) = (out (formula_reg_name j); out " := "; out_f f 0) | out_assign (AssignRelReg ((_, j), r)) = (out (rel_reg_name j); out " := "; out_r r 0) | out_assign (AssignIntReg (j, i)) = (out (int_reg_name j); out " := "; out_i i 0) and out_columns [] = () | out_columns [i] = out_i i 0 | out_columns (i :: is) = (out_i i 0; out ", "; out_columns is) and out_problem {comment, settings, univ_card, tuple_assigns, bounds, int_bounds, expr_assigns, formula} = (out ("\n" ^ block_comment comment ^ implode (map (fn (key, value) => key ^ ": " ^ value ^ "\n") settings) ^ "univ: " ^ atom_seq_name (univ_card, 0) ^ "\n" ^ implode (map string_for_tuple_assign tuple_assigns) ^ implode (map string_for_bound bounds) ^ (if int_bounds = [] then "" else "int_bounds: " ^ commas (map int_string_for_bound int_bounds) ^ "\n")); List.app (fn b => (out_assign b; out ";")) expr_assigns; out "solve "; out_outmost_f formula; out ";\n") in out ("// This file was generated by Isabelle (most likely Nitpick)\n" ^ "// " ^ ATP_Util.timestamp () ^ "\n"); List.app out_problem problems end (** Parsing of solution **) fun is_ident_char s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse s = "_" orelse s = "'" orelse s = "$" fun strip_blanks [] = [] | strip_blanks (" " :: ss) = strip_blanks ss | strip_blanks [s1, " "] = [s1] | strip_blanks (s1 :: " " :: s2 :: ss) = if is_ident_char s1 andalso is_ident_char s2 then s1 :: " " :: strip_blanks (s2 :: ss) else strip_blanks (s1 :: s2 :: ss) | strip_blanks (s :: ss) = s :: strip_blanks ss val scan_nat = Scan.repeat1 (Scan.one Symbol.is_ascii_digit) >> (the o Int.fromString o implode) val scan_rel_name = ($$ "s" |-- scan_nat >> pair 1 || $$ "r" |-- scan_nat >> pair 2 || ($$ "m" |-- scan_nat --| $$ "_") -- scan_nat) -- Scan.option ($$ "'") >> (fn ((n, j), SOME _) => (n, ~j - 1) | ((n, j), NONE) => (n, j)) val scan_atom = $$ "A" |-- scan_nat fun parse_non_empty_list scan = scan ::: Scan.repeat ($$ "," |-- scan) fun parse_list scan = parse_non_empty_list scan || Scan.succeed [] val parse_tuple = $$ "[" |-- parse_list scan_atom --| $$ "]" val parse_tuple_set = $$ "[" |-- parse_list parse_tuple --| $$ "]" val parse_assignment = (scan_rel_name --| $$ "=") -- parse_tuple_set val parse_instance = Scan.this_string "relations:" |-- $$ "{" |-- parse_list parse_assignment --| $$ "}" val extract_instance = fst o Scan.finite Symbol.stopper (Scan.error (!! (fn _ => raise SYNTAX ("Kodkod.extract_instance", "ill-formed Kodkodi output")) parse_instance)) o strip_blanks o raw_explode val problem_marker = "*** PROBLEM " val outcome_marker = "---OUTCOME---\n" val instance_marker = "---INSTANCE---\n" fun read_section_body marker = Substring.string o fst o Substring.position "\n\n" o Substring.triml (size marker) fun read_next_instance s = let val s = Substring.position instance_marker s |> snd in if Substring.isEmpty s then raise SYNTAX ("Kodkod.read_next_instance", "expected \"INSTANCE\" marker") else read_section_body instance_marker s |> extract_instance end fun read_next_outcomes j (s, ps, js) = let val (s1, s2) = Substring.position outcome_marker s in if Substring.isEmpty s2 orelse not (Substring.isEmpty (Substring.position problem_marker s1 |> snd)) then (s, ps, js) else let val outcome = read_section_body outcome_marker s2 val s = Substring.triml (size outcome_marker) s2 in if String.isSuffix "UNSATISFIABLE" outcome then read_next_outcomes j (s, ps, j :: js) else if String.isSuffix "SATISFIABLE" outcome then read_next_outcomes j (s, (j, read_next_instance s2) :: ps, js) else raise SYNTAX ("Kodkod.read_next_outcomes", "unknown outcome " ^ quote outcome) end end fun read_next_problems (s, ps, js) = let val s = Substring.position problem_marker s |> snd in if Substring.isEmpty s then (ps, js) else let val s = Substring.triml (size problem_marker) s val j_plus_1 = s |> Substring.takel (not_equal #" ") |> Substring.string |> Int.fromString |> the val j = j_plus_1 - 1 in read_next_problems (read_next_outcomes j (s, ps, js)) end end handle Option.Option => raise SYNTAX ("Kodkod.read_next_problems", "expected number after \"PROBLEM\"") (** Main Kodkod entry point **) fun serial_string_and_temporary_dir overlord = if overlord then ("", getenv "ISABELLE_HOME_USER") else (serial_string (), getenv "ISABELLE_TMP") (* The fudge term below is to account for Kodkodi's slow start-up time, which is partly due to the JVM and partly due to the ML "bash" function. *) val fudge_ms = 250 fun uncached_solve_any_problem overlord deadline max_threads0 max_solutions problems = let val j = find_index (curry (op =) True o #formula) problems val indexed_problems = if j >= 0 then [(j, nth problems j)] else filter_out (is_problem_trivially_false o snd) (0 upto length problems - 1 ~~ problems) val triv_js = filter_out (AList.defined (op =) indexed_problems) (0 upto length problems - 1) val reindex = fst o nth indexed_problems val max_threads = if max_threads0 = 0 then Options.default_int \<^system_option>\kodkod_max_threads\ else max_threads0 val external_process = not (Options.default_bool \<^system_option>\kodkod_scala\) orelse overlord val timeout0 = Time.toMilliseconds (deadline - Time.now ()) val timeout = if external_process then timeout0 - fudge_ms else timeout0 val solve_all = max_solutions > 1 in if null indexed_problems then Normal ([], triv_js, "") else if timeout <= 0 then TimedOut triv_js else let val kki = let val buf = Unsynchronized.ref Buffer.empty fun out s = Unsynchronized.change buf (Buffer.add s) val _ = write_problem out (map snd indexed_problems) in Buffer.content (! buf) end val (rc, out, err) = if external_process then let val (serial_str, temp_dir) = serial_string_and_temporary_dir overlord fun path_for suf = Path.explode (temp_dir ^ "/kodkodi" ^ serial_str ^ "." ^ suf) val kki_path = path_for "kki" val out_path = path_for "out" val err_path = path_for "err" fun remove_temporary_files () = if overlord then () else List.app (ignore o try File.rm) [kki_path, out_path, err_path] in let val _ = File.write kki_path kki val rc = Isabelle_System.bash ("cd " ^ Bash.string temp_dir ^ ";\n\ \\"$KODKODI/bin/kodkodi\"" ^ (" -max-msecs " ^ string_of_int timeout) ^ (if solve_all then " -solve-all" else "") ^ " -max-solutions " ^ string_of_int max_solutions ^ (if max_threads > 0 then " -max-threads " ^ string_of_int max_threads else "") ^ " < " ^ File.bash_path kki_path ^ " > " ^ File.bash_path out_path ^ " 2> " ^ File.bash_path err_path) val out = File.read out_path val err = File.read err_path val _ = remove_temporary_files () in (rc, out, err) end handle exn => (remove_temporary_files (); Exn.reraise exn) end else (timeout, (solve_all, (max_solutions, (max_threads, kki)))) |> let open XML.Encode in pair int (pair bool (pair int (pair int string))) end |> YXML.string_of_body |> \<^scala>\kodkod\ |> YXML.parse_body |> let open XML.Decode in triple int string string end val (ps, nontriv_js) = read_next_problems (Substring.full out, [], []) |>> rev ||> rev |> apfst (map (apfst reindex)) |> apsnd (map reindex) val js = triv_js @ nontriv_js val first_error = trim_split_lines err |> map (perhaps (try (unsuffix ".")) #> perhaps (try (unprefix "Solve error: ")) #> perhaps (try (unprefix "Error: "))) |> find_first (fn line => line <> "" andalso line <> "EXIT") |> the_default "" in if null ps then if rc = 130 then raise Exn.Interrupt else if rc = 2 then TimedOut js else if rc = 0 then Normal ([], js, first_error) - else if rc = 127 then - JavaNotFound else if first_error <> "" then Error (first_error, js) else Error ("Unknown error", js) else Normal (ps, js, first_error) end end val cached_outcome = Synchronized.var "Kodkod.cached_outcome" (NONE : ((int * problem list) * outcome) option) fun solve_any_problem debug overlord deadline max_threads max_solutions problems = let fun do_solve () = uncached_solve_any_problem overlord deadline max_threads max_solutions problems in if debug orelse overlord then do_solve () else case AList.lookup (fn ((max1, ps1), (max2, ps2)) => max1 = max2 andalso length ps1 = length ps2 andalso forall problems_equivalent (ps1 ~~ ps2)) (the_list (Synchronized.value cached_outcome)) (max_solutions, problems) of SOME outcome => outcome | NONE => let val outcome = do_solve () in (case outcome of Normal (_, _, "") => Synchronized.change cached_outcome (K (SOME ((max_solutions, problems), outcome))) | _ => ()); outcome end end end; diff --git a/src/HOL/Tools/Nitpick/nitpick.ML b/src/HOL/Tools/Nitpick/nitpick.ML --- a/src/HOL/Tools/Nitpick/nitpick.ML +++ b/src/HOL/Tools/Nitpick/nitpick.ML @@ -1,1011 +1,1003 @@ (* Title: HOL/Tools/Nitpick/nitpick.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2008, 2009, 2010 Finite model generation for HOL formulas using Kodkod. *) signature NITPICK = sig type term_postprocessor = Nitpick_Model.term_postprocessor datatype mode = Auto_Try | Try | Normal | TPTP type params = {cards_assigns: (typ option * int list) list, maxes_assigns: ((string * typ) option * int list) list, iters_assigns: ((string * typ) option * int list) list, bitss: int list, bisim_depths: int list, boxes: (typ option * bool option) list, finitizes: (typ option * bool option) list, monos: (typ option * bool option) list, wfs: ((string * typ) option * bool option) list, sat_solver: string, falsify: bool, debug: bool, verbose: bool, overlord: bool, spy: bool, user_axioms: bool option, assms: bool, whacks: term list, merge_type_vars: bool, binary_ints: bool option, destroy_constrs: bool, specialize: bool, star_linear_preds: bool, total_consts: bool option, needs: term list option, peephole_optim: bool, datatype_sym_break: int, kodkod_sym_break: int, timeout: Time.time, tac_timeout: Time.time, max_threads: int, show_types: bool, show_skolems: bool, show_consts: bool, evals: term list, formats: (term option * int list) list, atomss: (typ option * string list) list, max_potential: int, max_genuine: int, batch_size: int, expect: string} val genuineN : string val quasi_genuineN : string val potentialN : string val noneN : string val unknownN : string val register_frac_type : string -> (string * string) list -> theory -> theory val unregister_frac_type : string -> theory -> theory val register_codatatype : typ -> string -> (string * typ) list -> theory -> theory val unregister_codatatype : typ -> theory -> theory val register_term_postprocessor : typ -> term_postprocessor -> theory -> theory val unregister_term_postprocessor : typ -> theory -> theory val pick_nits_in_term : Proof.state -> params -> mode -> int -> int -> int -> (term * term) list -> term list -> term list -> term -> string * string list val pick_nits_in_subgoal : Proof.state -> params -> mode -> int -> int -> string * string list end; structure Nitpick : NITPICK = struct open Nitpick_Util open Nitpick_HOL open Nitpick_Preproc open Nitpick_Mono open Nitpick_Scope open Nitpick_Peephole open Nitpick_Rep open Nitpick_Nut open Nitpick_Kodkod open Nitpick_Model structure KK = Kodkod datatype mode = Auto_Try | Try | Normal | TPTP fun is_mode_nt mode = (mode = Normal orelse mode = TPTP) type params = {cards_assigns: (typ option * int list) list, maxes_assigns: ((string * typ) option * int list) list, iters_assigns: ((string * typ) option * int list) list, bitss: int list, bisim_depths: int list, boxes: (typ option * bool option) list, finitizes: (typ option * bool option) list, monos: (typ option * bool option) list, wfs: ((string * typ) option * bool option) list, sat_solver: string, falsify: bool, debug: bool, verbose: bool, overlord: bool, spy: bool, user_axioms: bool option, assms: bool, whacks: term list, merge_type_vars: bool, binary_ints: bool option, destroy_constrs: bool, specialize: bool, star_linear_preds: bool, total_consts: bool option, needs: term list option, peephole_optim: bool, datatype_sym_break: int, kodkod_sym_break: int, timeout: Time.time, tac_timeout: Time.time, max_threads: int, show_types: bool, show_skolems: bool, show_consts: bool, evals: term list, formats: (term option * int list) list, atomss: (typ option * string list) list, max_potential: int, max_genuine: int, batch_size: int, expect: string} val genuineN = "genuine" val quasi_genuineN = "quasi_genuine" val potentialN = "potential" val noneN = "none" val unknownN = "unknown" (* TODO: eliminate these historical aliases *) val register_frac_type = Nitpick_HOL.register_frac_type_global val unregister_frac_type = Nitpick_HOL.unregister_frac_type_global val register_codatatype = Nitpick_HOL.register_codatatype_global val unregister_codatatype = Nitpick_HOL.unregister_codatatype_global val register_term_postprocessor = Nitpick_Model.register_term_postprocessor_global val unregister_term_postprocessor = Nitpick_Model.unregister_term_postprocessor_global type problem_extension = {free_names: nut list, sel_names: nut list, nonsel_names: nut list, rel_table: nut NameTable.table, unsound: bool, scope: scope} type rich_problem = KK.problem * problem_extension fun pretties_for_formulas _ _ [] = [] | pretties_for_formulas ctxt s ts = [Pretty.str (s ^ plural_s_for_list ts ^ ":"), Pretty.indent indent_size (Pretty.chunks (map2 (fn j => fn t => Pretty.block [t |> shorten_names_in_term |> Syntax.pretty_term ctxt]) (length ts downto 1) ts))] -val isabelle_wrong_message = - "something appears to be wrong with your Isabelle installation" -val java_not_found_message = - "Java could not be launched -- " ^ isabelle_wrong_message - val max_unsound_delay_ms = 200 val max_unsound_delay_percent = 2 fun unsound_delay_for_timeout timeout = Int.max (0, Int.min (max_unsound_delay_ms, Time.toMilliseconds timeout * max_unsound_delay_percent div 100)) fun none_true assigns = forall (not_equal (SOME true) o snd) assigns fun has_lonely_bool_var (\<^const>\Pure.conjunction\ $ (\<^const>\Trueprop\ $ Free _) $ _) = true | has_lonely_bool_var _ = false val syntactic_sorts = \<^sort>\{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,equal}\ @ \<^sort>\numeral\ fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) = subset (op =) (S, syntactic_sorts) | has_tfree_syntactic_sort _ = false val has_syntactic_sorts = exists_type (exists_subtype has_tfree_syntactic_sort) fun plazy f = Pretty.para (f ()) fun pick_them_nits_in_term deadline state (params : params) mode i n step subst def_assm_ts nondef_assm_ts orig_t = let val time_start = Time.now () val timer = Timer.startRealTimer () val thy = Proof.theory_of state val ctxt = Proof.context_of state val keywords = Thy_Header.get_keywords thy val {cards_assigns, maxes_assigns, iters_assigns, bitss, bisim_depths, boxes, finitizes, monos, wfs, sat_solver, falsify, debug, verbose, overlord, spy, user_axioms, assms, whacks, merge_type_vars, binary_ints, destroy_constrs, specialize, star_linear_preds, total_consts, needs, peephole_optim, datatype_sym_break, kodkod_sym_break, tac_timeout, max_threads, show_types, show_skolems, show_consts, evals, formats, atomss, max_potential, max_genuine, batch_size, ...} = params val outcome = Synchronized.var "Nitpick.outcome" (Queue.empty: string Queue.T) fun pprint print prt = if mode = Auto_Try then Synchronized.change outcome (Queue.enqueue (Pretty.string_of prt)) else print (Pretty.string_of prt) val pprint_a = pprint writeln fun pprint_nt f = () |> is_mode_nt mode ? pprint writeln o f fun pprint_v f = () |> verbose ? pprint writeln o f fun pprint_d f = () |> debug ? pprint tracing o f val print = pprint writeln o Pretty.para fun print_t f = () |> mode = TPTP ? print o f (* val print_g = pprint tracing o Pretty.str *) val print_nt = pprint_nt o K o plazy val print_v = pprint_v o K o plazy fun check_deadline () = let val t = Time.now () in if Time.compare (t, deadline) <> LESS then raise Timeout.TIMEOUT (t - time_start) else () end val (def_assm_ts, nondef_assm_ts) = if assms orelse mode <> Normal then (def_assm_ts, nondef_assm_ts) else ([], []) val _ = if step = 0 then print_nt (fn () => "Nitpicking formula...") else pprint_nt (fn () => Pretty.chunks ( pretties_for_formulas ctxt ("Nitpicking " ^ (if i <> 1 orelse n <> 1 then "subgoal " ^ string_of_int i ^ " of " ^ string_of_int n else "goal")) [Logic.list_implies (nondef_assm_ts, orig_t)])) val _ = spying spy (fn () => (state, i, "starting " ^ \<^make_string> mode ^ " mode")) val _ = print_v (prefix "Timestamp: " o Date.fmt "%H:%M:%S" o Date.fromTimeLocal o Time.now) val neg_t = if falsify then Logic.mk_implies (orig_t, \<^prop>\False\) else orig_t val conj_ts = neg_t :: def_assm_ts @ nondef_assm_ts @ evals @ these needs val tfree_table = if merge_type_vars then merged_type_var_table_for_terms thy conj_ts else [] val merge_tfrees = merge_type_vars_in_term thy merge_type_vars tfree_table val neg_t = neg_t |> merge_tfrees val def_assm_ts = def_assm_ts |> map merge_tfrees val nondef_assm_ts = nondef_assm_ts |> map merge_tfrees val evals = evals |> map merge_tfrees val needs = needs |> Option.map (map merge_tfrees) val conj_ts = neg_t :: def_assm_ts @ nondef_assm_ts @ evals @ these needs val original_max_potential = max_potential val original_max_genuine = max_genuine val max_bisim_depth = fold Integer.max bisim_depths ~1 val case_names = case_const_names ctxt val defs = def_assm_ts @ all_defs_of thy subst val nondefs = all_nondefs_of ctxt subst val def_tables = const_def_tables ctxt subst defs val nondef_table = const_nondef_table nondefs val simp_table = Unsynchronized.ref (const_simp_table ctxt subst) val psimp_table = const_psimp_table ctxt subst val choice_spec_table = const_choice_spec_table ctxt subst val nondefs = nondefs |> filter_out (is_choice_spec_axiom ctxt choice_spec_table) val intro_table = inductive_intro_table ctxt subst def_tables val ground_thm_table = ground_theorem_table thy val ersatz_table = ersatz_table ctxt val hol_ctxt as {wf_cache, ...} = {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes, wfs = wfs, user_axioms = user_axioms, debug = debug, whacks = whacks, binary_ints = binary_ints, destroy_constrs = destroy_constrs, specialize = specialize, star_linear_preds = star_linear_preds, total_consts = total_consts, needs = needs, tac_timeout = tac_timeout, evals = evals, case_names = case_names, def_tables = def_tables, nondef_table = nondef_table, nondefs = nondefs, simp_table = simp_table, psimp_table = psimp_table, choice_spec_table = choice_spec_table, intro_table = intro_table, ground_thm_table = ground_thm_table, ersatz_table = ersatz_table, skolems = Unsynchronized.ref [], special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []} val pseudo_frees = [] val real_frees = fold Term.add_frees conj_ts [] val _ = null (fold Term.add_tvars conj_ts []) orelse error "Nitpick cannot handle goals with schematic type variables" val (nondef_ts, def_ts, need_ts, got_all_mono_user_axioms, no_poly_user_axioms, binarize) = preprocess_formulas hol_ctxt nondef_assm_ts neg_t val got_all_user_axioms = got_all_mono_user_axioms andalso no_poly_user_axioms fun print_wf (x, (gfp, wf)) = pprint_nt (fn () => Pretty.blk (0, Pretty.text ("The " ^ (if gfp then "co" else "") ^ "inductive predicate") @ [Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (Const x)), Pretty.brk 1] @ Pretty.text (if wf then "was proved well-founded; Nitpick can compute it \ \efficiently" else "could not be proved well-founded; Nitpick might need to \ \unroll it"))) val _ = if verbose then List.app print_wf (!wf_cache) else () val das_wort_formula = if falsify then "Negated conjecture" else "Formula" val _ = pprint_d (fn () => Pretty.chunks (pretties_for_formulas ctxt das_wort_formula [hd nondef_ts] @ pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @ pretties_for_formulas ctxt "Relevant nondefinitional axiom" (tl nondef_ts))) val _ = List.app (ignore o Term.type_of) (nondef_ts @ def_ts) handle TYPE (_, Ts, ts) => raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts) val nondef_us = nondef_ts |> map (nut_from_term hol_ctxt Eq) val def_us = def_ts |> map (nut_from_term hol_ctxt DefEq) val need_us = need_ts |> map (nut_from_term hol_ctxt Eq) val (free_names, const_names) = fold add_free_and_const_names (nondef_us @ def_us @ need_us) ([], []) val (sel_names, nonsel_names) = List.partition (is_sel o nickname_of) const_names val sound_finitizes = none_true finitizes (* val _ = List.app (print_g o string_for_nut ctxt) (nondef_us @ def_us @ need_us) *) val unique_scope = forall (curry (op =) 1 o length o snd) cards_assigns fun monotonicity_message Ts extra = let val pretties = map (pretty_maybe_quote keywords o pretty_for_type ctxt) Ts in Pretty.blk (0, Pretty.text ("The type" ^ plural_s_for_list pretties) @ [Pretty.brk 1] @ pretty_serial_commas "and" pretties @ [Pretty.brk 1] @ Pretty.text ( (if none_true monos then "passed the monotonicity test" else (if length pretties = 1 then "is" else "are") ^ " considered monotonic") ^ "; " ^ extra)) end fun is_type_fundamentally_monotonic T = (is_data_type ctxt T andalso not (is_quot_type ctxt T) andalso (not (is_pure_typedef ctxt T) orelse is_univ_typedef ctxt T)) orelse is_number_type ctxt T orelse is_bit_type T fun is_type_actually_monotonic T = Timeout.apply tac_timeout (formulas_monotonic hol_ctxt binarize T) (nondef_ts, def_ts) handle Timeout.TIMEOUT _ => false fun is_type_kind_of_monotonic T = case triple_lookup (type_match thy) monos T of SOME (SOME false) => false | _ => is_type_actually_monotonic T fun is_type_monotonic T = unique_scope orelse case triple_lookup (type_match thy) monos T of SOME (SOME b) => b | _ => is_type_fundamentally_monotonic T orelse is_type_actually_monotonic T fun is_deep_data_type_finitizable T = triple_lookup (type_match thy) finitizes T = SOME (SOME true) fun is_shallow_data_type_finitizable (T as Type (\<^type_name>\fun_box\, _)) = is_type_kind_of_monotonic T | is_shallow_data_type_finitizable T = case triple_lookup (type_match thy) finitizes T of SOME (SOME b) => b | _ => is_type_kind_of_monotonic T fun is_data_type_deep T = T = \<^typ>\unit\ orelse T = nat_T orelse is_word_type T orelse exists (curry (op =) T o domain_type o type_of) sel_names val all_Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts) |> sort Term_Ord.typ_ord val _ = if verbose andalso binary_ints = SOME true andalso exists (member (op =) [nat_T, int_T]) all_Ts then print_v (K "The option \"binary_ints\" will be ignored because of the \ \presence of rationals, reals, \"Suc\", \"gcd\", or \"lcm\" \ \in the problem") else () val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts val _ = if verbose andalso not unique_scope then case filter_out is_type_fundamentally_monotonic mono_Ts of [] => () | interesting_mono_Ts => pprint_v (K (monotonicity_message interesting_mono_Ts "Nitpick might be able to skip some scopes")) else () val (deep_dataTs, shallow_dataTs) = all_Ts |> filter (is_data_type ctxt) |> List.partition is_data_type_deep val finitizable_dataTs = (deep_dataTs |> filter_out (is_finite_type hol_ctxt) |> filter is_deep_data_type_finitizable) @ (shallow_dataTs |> filter_out (is_finite_type hol_ctxt) |> filter is_shallow_data_type_finitizable) val _ = if verbose andalso not (null finitizable_dataTs) then pprint_v (K (monotonicity_message finitizable_dataTs "Nitpick can use a more precise finite encoding")) else () (* val _ = print_g "Monotonic types:" val _ = List.app (print_g o string_for_type ctxt) mono_Ts val _ = print_g "Nonmonotonic types:" val _ = List.app (print_g o string_for_type ctxt) nonmono_Ts *) val incremental = Int.max (max_potential, max_genuine) >= 2 val actual_sat_solver = if sat_solver <> "smart" then if incremental andalso not (member (op =) (Kodkod_SAT.configured_sat_solvers true) sat_solver) then (print_nt (K ("An incremental SAT solver is required: \"SAT4J\" will \ \be used instead of " ^ quote sat_solver)); "SAT4J") else sat_solver else Kodkod_SAT.smart_sat_solver_name incremental val _ = if sat_solver = "smart" then print_v (fn () => "Using SAT solver " ^ quote actual_sat_solver ^ "\nThe following" ^ (if incremental then " incremental " else " ") ^ "solvers are configured: " ^ commas_quote (Kodkod_SAT.configured_sat_solvers incremental)) else () val too_big_scopes = Unsynchronized.ref [] fun problem_for_scope unsound (scope as {card_assigns, bits, bisim_depth, data_types, ofs, ...}) = let val _ = not (exists (fn other => scope_less_eq other scope) (!too_big_scopes)) orelse raise TOO_LARGE ("Nitpick.pick_them_nits_in_term.\ \problem_for_scope", "too large scope") (* val _ = print_g "Offsets:" val _ = List.app (fn (T, j0) => print_g (string_for_type ctxt T ^ " = " ^ string_of_int j0)) (Typtab.dest ofs) *) val effective_total_consts = case total_consts of SOME b => b | NONE => forall (is_exact_type data_types true) all_Ts val main_j0 = offset_of_type ofs bool_T val (nat_card, nat_j0) = spec_of_type scope nat_T val (int_card, int_j0) = spec_of_type scope int_T val _ = (nat_j0 = main_j0 andalso int_j0 = main_j0) orelse raise BAD ("Nitpick.pick_them_nits_in_term.problem_for_scope", "bad offsets") val kk = kodkod_constrs peephole_optim nat_card int_card main_j0 val (free_names, rep_table) = choose_reps_for_free_vars scope pseudo_frees free_names NameTable.empty val (sel_names, rep_table) = choose_reps_for_all_sels scope rep_table val (nonsel_names, rep_table) = choose_reps_for_consts scope effective_total_consts nonsel_names rep_table val (guiltiest_party, min_highest_arity) = NameTable.fold (fn (name, R) => fn (s, n) => let val n' = arity_of_rep R in if n' > n then (nickname_of name, n') else (s, n) end) rep_table ("", 1) val min_univ_card = NameTable.fold (Integer.max o min_univ_card_of_rep o snd) rep_table (univ_card nat_card int_card main_j0 [] KK.True) val _ = check_arity guiltiest_party min_univ_card min_highest_arity val def_us = def_us |> map (choose_reps_in_nut scope unsound rep_table true) val nondef_us = nondef_us |> map (choose_reps_in_nut scope unsound rep_table false) val need_us = need_us |> map (choose_reps_in_nut scope unsound rep_table false) (* val _ = List.app (print_g o string_for_nut ctxt) (free_names @ sel_names @ nonsel_names @ nondef_us @ def_us @ need_us) *) val (free_rels, pool, rel_table) = rename_free_vars free_names initial_pool NameTable.empty val (sel_rels, pool, rel_table) = rename_free_vars sel_names pool rel_table val (other_rels, pool, rel_table) = rename_free_vars nonsel_names pool rel_table val nondef_us = nondef_us |> map (rename_vars_in_nut pool rel_table) val def_us = def_us |> map (rename_vars_in_nut pool rel_table) val need_us = need_us |> map (rename_vars_in_nut pool rel_table) val nondef_fs = map (kodkod_formula_from_nut ofs kk) nondef_us val def_fs = map (kodkod_formula_from_nut ofs kk) def_us val formula = fold (fold s_and) [def_fs, nondef_fs] KK.True val comment = (if unsound then "unsound" else "sound") ^ "\n" ^ multiline_string_for_scope scope val kodkod_sat_solver = Kodkod_SAT.sat_solver_spec actual_sat_solver |> snd val bit_width = if bits = 0 then 16 else bits + 1 val delay = if unsound then unsound_delay_for_timeout (deadline - Time.now ()) else 0 val settings = [("solver", commas_quote kodkod_sat_solver), ("bit_width", string_of_int bit_width), ("symmetry_breaking", string_of_int kodkod_sym_break), ("sharing", "3"), ("flatten", "false"), ("delay", signed_string_of_int delay)] val plain_rels = free_rels @ other_rels val plain_bounds = map (bound_for_plain_rel ctxt debug) plain_rels val plain_axioms = map (declarative_axiom_for_plain_rel kk) plain_rels val need_vals = map (fn dtype as {typ, ...} => (typ, needed_values_for_data_type need_us ofs dtype)) data_types val sel_bounds = map (bound_for_sel_rel ctxt debug need_vals data_types) sel_rels val dtype_axioms = declarative_axioms_for_data_types hol_ctxt binarize need_us need_vals datatype_sym_break bits ofs kk rel_table data_types val declarative_axioms = plain_axioms @ dtype_axioms val univ_card = Int.max (univ_card nat_card int_card main_j0 (plain_bounds @ sel_bounds) formula, main_j0 |> bits > 0 ? Integer.add (bits + 1)) val (built_in_bounds, built_in_axioms) = bounds_and_axioms_for_built_in_rels_in_formulas debug univ_card nat_card int_card main_j0 (formula :: declarative_axioms) val bounds = built_in_bounds @ plain_bounds @ sel_bounds |> not debug ? merge_bounds val axioms = built_in_axioms @ declarative_axioms val highest_arity = fold Integer.max (map (fst o fst) (maps fst bounds)) 0 val formula = fold_rev s_and axioms formula val _ = if bits = 0 then () else check_bits bits formula val _ = if formula = KK.False then () else check_arity "" univ_card highest_arity in SOME ({comment = comment, settings = settings, univ_card = univ_card, tuple_assigns = [], bounds = bounds, int_bounds = if bits = 0 then sequential_int_bounds univ_card else pow_of_two_int_bounds bits main_j0, expr_assigns = [], formula = formula}, {free_names = free_names, sel_names = sel_names, nonsel_names = nonsel_names, rel_table = rel_table, unsound = unsound, scope = scope}) end handle TOO_LARGE (loc, msg) => if loc = "Nitpick_Kodkod.check_arity" andalso not (Typtab.is_empty ofs) then problem_for_scope unsound {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns, bits = bits, bisim_depth = bisim_depth, data_types = data_types, ofs = Typtab.empty} else if loc = "Nitpick.pick_them_nits_in_term.\ \problem_for_scope" then NONE else (Unsynchronized.change too_big_scopes (cons scope); print_v (fn () => "Limit reached: " ^ msg ^ "; skipping " ^ (if unsound then "potential component of " else "") ^ "scope"); NONE) | TOO_SMALL (_, msg) => (print_v (fn () => "Limit reached: " ^ msg ^ "; skipping " ^ (if unsound then "potential component of " else "") ^ "scope"); NONE) val das_wort_model = if falsify then "counterexample" else "model" val scopes = Unsynchronized.ref [] val generated_scopes = Unsynchronized.ref [] val generated_problems = Unsynchronized.ref ([] : rich_problem list) val checked_problems = Unsynchronized.ref (SOME []) val met_potential = Unsynchronized.ref 0 fun update_checked_problems problems = List.app (Unsynchronized.change checked_problems o Option.map o cons o nth problems) fun show_kodkod_warning "" = () | show_kodkod_warning s = print_nt (fn () => "Kodkod warning: " ^ s) fun print_and_check_model genuine bounds ({free_names, sel_names, nonsel_names, rel_table, scope, ...} : problem_extension) = let val _ = print_t (fn () => "% SZS status " ^ (if genuine then if falsify then "CounterSatisfiable" else "Satisfiable" else "GaveUp") ^ "\n" ^ "% SZS output start FiniteModel") val (reconstructed_model, codatatypes_ok) = reconstruct_hol_model {show_types = show_types, show_skolems = show_skolems, show_consts = show_consts} scope formats atomss real_frees pseudo_frees free_names sel_names nonsel_names rel_table bounds val genuine_means_genuine = got_all_user_axioms andalso none_true wfs andalso sound_finitizes andalso total_consts <> SOME true andalso codatatypes_ok fun assms_prop () = Logic.mk_conjunction_list (neg_t :: def_assm_ts @ nondef_assm_ts) in (pprint_a (Pretty.chunks [Pretty.blk (0, (Pretty.text ((if mode = Auto_Try then "Auto " else "") ^ "Nitpick found a" ^ (if not genuine then " potentially spurious " else if genuine_means_genuine then " " else " quasi genuine ") ^ das_wort_model) @ (case pretties_for_scope scope verbose of [] => [] | pretties => [Pretty.brk 1, Pretty.str "for", Pretty.brk 1] @ pretties) @ [Pretty.str ":"])), Pretty.indent indent_size reconstructed_model]); print_t (K "% SZS output end FiniteModel"); if genuine then (if has_lonely_bool_var orig_t then print "Hint: Maybe you forgot a colon after the lemma's name?" else if has_syntactic_sorts orig_t then print "Hint: Maybe you forgot a type constraint?" else (); if not genuine_means_genuine then if no_poly_user_axioms then let val options = [] |> not got_all_mono_user_axioms ? cons ("user_axioms", "\"true\"") |> not (none_true wfs) ? cons ("wf", "\"smart\" or \"false\"") |> not sound_finitizes ? cons ("finitize", "\"smart\" or \"false\"") |> total_consts = SOME true ? cons ("total_consts", "\"smart\" or \"false\"") |> not codatatypes_ok ? cons ("bisim_depth", "a nonnegative value") val ss = map (fn (name, value) => quote name ^ " set to " ^ value) options in print ("Try again with " ^ space_implode " " (serial_commas "and" ss) ^ " to confirm that the " ^ das_wort_model ^ " is genuine") end else print ("Nitpick is unable to guarantee the authenticity of \ \the " ^ das_wort_model ^ " in the presence of \ \polymorphic axioms") else (); NONE) else if not genuine then (Unsynchronized.inc met_potential; NONE) else NONE) |> pair genuine_means_genuine end fun solve_any_problem (found_really_genuine, max_potential, max_genuine, donno) first_time problems = let val max_potential = Int.max (0, max_potential) val max_genuine = Int.max (0, max_genuine) fun print_and_check genuine (j, bounds) = print_and_check_model genuine bounds (snd (nth problems j)) val max_solutions = max_potential + max_genuine |> not incremental ? Integer.min 1 in if max_solutions <= 0 then (found_really_genuine, 0, 0, donno) else case KK.solve_any_problem debug overlord deadline max_threads max_solutions (map fst problems) of - KK.JavaNotFound => - (print_nt (K java_not_found_message); - (found_really_genuine, max_potential, max_genuine, donno + 1)) - | KK.Normal ([], unsat_js, s) => + KK.Normal ([], unsat_js, s) => (update_checked_problems problems unsat_js; show_kodkod_warning s; (found_really_genuine, max_potential, max_genuine, donno)) | KK.Normal (sat_ps, unsat_js, s) => let val (lib_ps, con_ps) = List.partition (#unsound o snd o nth problems o fst) sat_ps in update_checked_problems problems (unsat_js @ map fst lib_ps); show_kodkod_warning s; if null con_ps then let val genuine_codes = lib_ps |> take max_potential |> map (print_and_check false) |> filter (curry (op =) (SOME true) o snd) val found_really_genuine = found_really_genuine orelse exists fst genuine_codes val num_genuine = length genuine_codes val max_genuine = max_genuine - num_genuine val max_potential = max_potential - (length lib_ps - num_genuine) in if max_genuine <= 0 then (found_really_genuine, 0, 0, donno) else let (* "co_js" is the list of sound problems whose unsound pendants couldn't be satisfied and hence that most probably can't be satisfied themselves. *) val co_js = map (fn j => j - 1) unsat_js |> filter (fn j => j >= 0 andalso scopes_equivalent (#scope (snd (nth problems j)), #scope (snd (nth problems (j + 1))))) val bye_js = sort_distinct int_ord (map fst sat_ps @ unsat_js @ co_js) val problems = problems |> filter_out_indices bye_js |> max_potential <= 0 ? filter_out (#unsound o snd) in solve_any_problem (found_really_genuine, max_potential, max_genuine, donno) false problems end end else let val genuine_codes = con_ps |> take max_genuine |> map (print_and_check true) val max_genuine = max_genuine - length genuine_codes val found_really_genuine = found_really_genuine orelse exists fst genuine_codes in if max_genuine <= 0 orelse not first_time then (found_really_genuine, 0, max_genuine, donno) else let val bye_js = sort_distinct int_ord (map fst sat_ps @ unsat_js) val problems = problems |> filter_out_indices bye_js |> filter_out (#unsound o snd) in solve_any_problem (found_really_genuine, 0, max_genuine, donno) false problems end end end | KK.TimedOut unsat_js => (update_checked_problems problems unsat_js; raise Timeout.TIMEOUT Time.zeroTime) | KK.Error (s, unsat_js) => (update_checked_problems problems unsat_js; print_v (K ("Kodkod error: " ^ s)); (found_really_genuine, max_potential, max_genuine, donno + 1)) end fun run_batch j n scopes (found_really_genuine, max_potential, max_genuine, donno) = let val _ = if null scopes then print_nt (K "The scope specification is inconsistent") else if verbose then pprint_nt (fn () => Pretty.chunks [Pretty.blk (0, Pretty.text ((if n > 1 then "Batch " ^ string_of_int (j + 1) ^ " of " ^ signed_string_of_int n ^ ": " else "") ^ "Trying " ^ string_of_int (length scopes) ^ " scope" ^ plural_s_for_list scopes ^ ":")), Pretty.indent indent_size (Pretty.chunks (map2 (fn j => fn scope => Pretty.block ( (case pretties_for_scope scope true of [] => [Pretty.str "Empty"] | pretties => pretties))) (length scopes downto 1) scopes))]) else () fun add_problem_for_scope (scope, unsound) (problems, donno) = (check_deadline (); case problem_for_scope unsound scope of SOME problem => (problems |> (null problems orelse not (KK.problems_equivalent (fst problem, fst (hd problems)))) ? cons problem, donno) | NONE => (problems, donno + 1)) val (problems, donno) = fold add_problem_for_scope (map_product pair scopes ((if max_genuine > 0 then [false] else []) @ (if max_potential > 0 then [true] else []))) ([], donno) val _ = Unsynchronized.change generated_problems (append problems) val _ = Unsynchronized.change generated_scopes (append scopes) val _ = if j + 1 = n then let val (unsound_problems, sound_problems) = List.partition (#unsound o snd) (!generated_problems) in if not (null sound_problems) andalso forall (KK.is_problem_trivially_false o fst) sound_problems then print_nt (fn () => "Warning: The conjecture " ^ (if falsify then "either trivially holds" else "is either trivially unsatisfiable") ^ " for the given scopes or lies outside Nitpick's supported \ \fragment" ^ (if exists (not o KK.is_problem_trivially_false o fst) unsound_problems then "; only potentially spurious " ^ das_wort_model ^ "s may be found" else "")) else () end else () in solve_any_problem (found_really_genuine, max_potential, max_genuine, donno) true (rev problems) end fun scope_count (problems : rich_problem list) scope = length (filter (curry scopes_equivalent scope o #scope o snd) problems) fun excipit did_so_and_so = let val do_filter = if !met_potential = max_potential then filter_out (#unsound o snd) else I val total = length (!scopes) val unsat = fold (fn scope => case scope_count (do_filter (!generated_problems)) scope of 0 => I | n => scope_count (do_filter (these (!checked_problems))) scope = n ? Integer.add 1) (!generated_scopes) 0 in (if mode = TPTP then "% SZS status GaveUp\n" else "") ^ "Nitpick " ^ did_so_and_so ^ (if is_some (!checked_problems) andalso total > 0 then " " ^ string_of_int unsat ^ " of " ^ string_of_int total ^ " scope" ^ plural_s total else "") end val (skipped, the_scopes) = all_scopes hol_ctxt binarize cards_assigns maxes_assigns iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs finitizable_dataTs val _ = if skipped > 0 then print_nt (fn () => "Skipping " ^ string_of_int skipped ^ " scope" ^ plural_s skipped ^ ". (Consider using \"mono\" or \ \\"merge_type_vars\" to prevent this.)") else () val _ = scopes := the_scopes fun run_batches _ _ [] (found_really_genuine, max_potential, max_genuine, donno) = if donno > 0 andalso max_genuine > 0 then (print_nt (fn () => excipit "checked"); unknownN) else if max_genuine = original_max_genuine then if max_potential = original_max_potential then (print_t (K "% SZS status GaveUp"); print_nt (fn () => "Nitpick found no " ^ das_wort_model); if skipped > 0 then unknownN else noneN) else (print_nt (fn () => excipit ("could not find a" ^ (if max_genuine = 1 then " better " ^ das_wort_model else "ny better " ^ das_wort_model ^ "s") ^ "\nIt checked")); potentialN) else if found_really_genuine then genuineN else quasi_genuineN | run_batches j n (batch :: batches) z = let val (z as (_, _, max_genuine, _)) = run_batch j n batch z in run_batches (j + 1) n (if max_genuine > 0 then batches else []) z end val batches = chunk_list batch_size (!scopes) val outcome_code = run_batches 0 (length batches) batches (false, max_potential, max_genuine, 0) handle Timeout.TIMEOUT _ => (print_nt (fn () => excipit "ran out of time after checking"); if !met_potential > 0 then potentialN else unknownN) val _ = print_v (fn () => "Total time: " ^ string_of_time (Timer.checkRealTimer timer)) val _ = spying spy (fn () => (state, i, "outcome: " ^ outcome_code)) in (outcome_code, Queue.content (Synchronized.value outcome)) end (* Give the inner timeout a chance. *) val timeout_bonus = seconds 1.0 fun pick_nits_in_term state (params as {timeout, expect, ...}) mode i n step subst def_assm_ts nondef_assm_ts orig_t = let val print_nt = if is_mode_nt mode then writeln else K () val print_t = if mode = TPTP then writeln else K () in let val unknown_outcome = (unknownN, []) val deadline = Time.now () + timeout val outcome as (outcome_code, _) = Timeout.apply (timeout + timeout_bonus) (pick_them_nits_in_term deadline state params mode i n step subst def_assm_ts nondef_assm_ts) orig_t handle TOO_LARGE (_, details) => (print_t "% SZS status GaveUp"; print_nt ("Limit reached: " ^ details); unknown_outcome) | TOO_SMALL (_, details) => (print_t "% SZS status GaveUp"; print_nt ("Limit reached: " ^ details); unknown_outcome) | Kodkod.SYNTAX (_, details) => (print_t "% SZS status GaveUp"; print_nt ("Malformed Kodkodi output: " ^ details); unknown_outcome) | Timeout.TIMEOUT _ => (print_t "% SZS status TimedOut"; print_nt "Nitpick ran out of time"; unknown_outcome) in if expect = "" orelse outcome_code = expect then outcome else error ("Unexpected outcome: " ^ quote outcome_code) end end fun is_fixed_equation ctxt (Const (\<^const_name>\Pure.eq\, _) $ Free (s, _) $ Const _) = Variable.is_fixed ctxt s | is_fixed_equation _ _ = false fun extract_fixed_frees ctxt (assms, t) = let val (subst, other_assms) = List.partition (is_fixed_equation ctxt) assms |>> map Logic.dest_equals in (subst, other_assms, subst_atomic subst t) end fun pick_nits_in_subgoal state params mode i step = let val ctxt = Proof.context_of state val t = state |> Proof.raw_goal |> #goal |> Thm.prop_of in case Logic.count_prems t of 0 => (writeln "No subgoal!"; (noneN, [])) | n => let val t = Logic.goal_params t i |> fst val assms = map Thm.term_of (Assumption.all_assms_of ctxt) val (subst, assms, t) = extract_fixed_frees ctxt (assms, t) in pick_nits_in_term state params mode i n step subst [] assms t end end end;