--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Dec 06 16:18:56 2010 +0100
@@ -23,23 +23,27 @@
open Nitpick_Util
open Nitpick_HOL
-type var = int
+structure PL = PropLogic
datatype sign = Plus | Minus
-datatype sign_atom = S of sign | V of var
+
+type var = int
-type literal = var * sign
+datatype annotation = Gen | New | Fls | Tru
+datatype annotation_atom = A of annotation | V of var
+
+type assign_literal = var * (sign * annotation)
datatype mtyp =
MAlpha |
- MFun of mtyp * sign_atom * mtyp |
+ MFun of mtyp * annotation_atom * mtyp |
MPair of mtyp * mtyp |
MType of string * mtyp list |
MRec of string * typ list
datatype mterm =
MRaw of term * mtyp |
- MAbs of string * typ * mtyp * sign_atom * mterm |
+ MAbs of string * typ * mtyp * annotation_atom * mterm |
MApp of mterm * mterm
type mdata =
@@ -58,22 +62,29 @@
val trace = Unsynchronized.ref false
fun trace_msg msg = if !trace then tracing (msg ()) else ()
+fun string_for_sign Plus = "+"
+ | string_for_sign Minus = "-"
+
+fun negate_sign Plus = Minus
+ | negate_sign Minus = Plus
+
val string_for_var = signed_string_of_int
fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>"
| string_for_vars sep xs = space_implode sep (map string_for_var xs)
fun subscript_string_for_vars sep xs =
if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>"
-fun string_for_sign Plus = "+"
- | string_for_sign Minus = "-"
+fun string_for_annotation Gen = "G"
+ | string_for_annotation New = "N"
+ | string_for_annotation Fls = "F"
+ | string_for_annotation Tru = "T"
-fun xor sn1 sn2 = if sn1 = sn2 then Plus else Minus
-val negate = xor Minus
+fun string_for_annotation_atom (A a) = string_for_annotation a
+ | string_for_annotation_atom (V x) = string_for_var x
-fun string_for_sign_atom (S sn) = string_for_sign sn
- | string_for_sign_atom (V x) = string_for_var x
-
-fun string_for_literal (x, sn) = string_for_var x ^ " = " ^ string_for_sign sn
+fun string_for_assign_literal (x, (sn, a)) =
+ string_for_var x ^ (case sn of Plus => " = " | Minus => " \<noteq> ") ^
+ string_for_annotation a
val bool_M = MType (@{type_name bool}, [])
val dummy_M = MType (nitpick_prefix ^ "dummy", [])
@@ -101,9 +112,9 @@
"_"
else case M of
MAlpha => "\<alpha>"
- | MFun (M1, a, M2) =>
+ | MFun (M1, aa, M2) =>
aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^
- string_for_sign_atom a ^ "\<^esup> " ^ aux prec M2
+ string_for_annotation_atom aa ^ "\<^esup> " ^ aux prec M2
| MPair (M1, M2) => aux (prec + 1) M1 ^ " \<times> " ^ aux prec M2
| MType (s, []) =>
if s = @{type_name prop} orelse s = @{type_name bool} then "o"
@@ -133,16 +144,16 @@
(if need_parens then "(" else "") ^
(case m of
MRaw (t, M) => Syntax.string_of_term ctxt t ^ mtype_annotation M
- | MAbs (s, _, M, a, m) =>
+ | MAbs (s, _, M, aa, m) =>
"\<lambda>" ^ s ^ mtype_annotation M ^ ".\<^bsup>" ^
- string_for_sign_atom a ^ "\<^esup> " ^ aux prec m
+ string_for_annotation_atom aa ^ "\<^esup> " ^ aux prec m
| MApp (m1, m2) => aux prec m1 ^ " " ^ aux (prec + 1) m2) ^
(if need_parens then ")" else "")
end
in aux 0 end
fun mtype_of_mterm (MRaw (_, M)) = M
- | mtype_of_mterm (MAbs (_, _, M, a, m)) = MFun (M, a, mtype_of_mterm m)
+ | mtype_of_mterm (MAbs (_, _, M, aa, m)) = MFun (M, aa, mtype_of_mterm m)
| mtype_of_mterm (MApp (m1, _)) =
case mtype_of_mterm m1 of
MFun (_, _, M12) => M12
@@ -185,11 +196,11 @@
| exists_alpha_sub_mtype_fresh (MRec _) = true
fun constr_mtype_for_binders z Ms =
- fold_rev (fn M => curry3 MFun M (S Minus)) Ms (MRec z)
+ fold_rev (fn M => curry3 MFun M (A Gen)) Ms (MRec z)
fun repair_mtype _ _ MAlpha = MAlpha
- | repair_mtype cache seen (MFun (M1, a, M2)) =
- MFun (repair_mtype cache seen M1, a, repair_mtype cache seen M2)
+ | repair_mtype cache seen (MFun (M1, aa, M2)) =
+ MFun (repair_mtype cache seen M1, aa, repair_mtype cache seen M2)
| repair_mtype cache seen (MPair Mp) =
MPair (pairself (repair_mtype cache seen) Mp)
| repair_mtype cache seen (MType (s, Ms)) =
@@ -218,6 +229,8 @@
| is_fin_fun_supported_type @{typ bool} = true
| is_fin_fun_supported_type (Type (@{type_name option}, _)) = true
| is_fin_fun_supported_type _ = false
+
+(* TODO: clean this up *)
fun fin_fun_body _ _ (t as @{term False}) = SOME t
| fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t
| fin_fun_body dom_T ran_T
@@ -233,17 +246,19 @@
$ (Const (@{const_name unknown}, ran_T)) $ (t0 $ t1 $ t2 $ t3)))
| fin_fun_body _ _ _ = NONE
+(* ### FIXME: make sure well-annotated! *)
+
fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus
T1 T2 =
let
val M1 = fresh_mtype_for_type mdata all_minus T1
val M2 = fresh_mtype_for_type mdata all_minus T2
- val a = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
- is_fin_fun_supported_type (body_type T2) then
- V (Unsynchronized.inc max_fresh)
- else
- S Minus
- in (M1, a, M2) end
+ val aa = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
+ is_fin_fun_supported_type (body_type T2) then
+ V (Unsynchronized.inc max_fresh)
+ else
+ A Gen
+ in (M1, aa, M2) end
and fresh_mtype_for_type (mdata as {hol_ctxt as {ctxt, ...}, binarize, alpha_T,
datatype_mcache, constr_mcache, ...})
all_minus =
@@ -294,14 +309,23 @@
| _ => MType (simple_string_of_typ T, [])
in do_type end
+val ground_and_sole_base_constrs = [] (* FIXME: [@{const_name Nil}, @{const_name None}], cf. lists_empty *)
+
fun prodM_factors (MPair (M1, M2)) = maps prodM_factors [M1, M2]
| prodM_factors M = [M]
fun curried_strip_mtype (MFun (M1, _, M2)) =
curried_strip_mtype M2 |>> append (prodM_factors M1)
| curried_strip_mtype M = ([], M)
fun sel_mtype_from_constr_mtype s M =
- let val (arg_Ms, dataM) = curried_strip_mtype M in
- MFun (dataM, S Minus,
+ let
+ val (arg_Ms, dataM) = curried_strip_mtype M
+ val a = if member (op =) ground_and_sole_base_constrs
+ (constr_name_for_sel_like s) then
+ Fls
+ else
+ Gen
+ in
+ MFun (dataM, A a,
case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n)
end
@@ -323,230 +347,437 @@
x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
|> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s
-fun resolve_sign_atom lits (V x) =
- x |> AList.lookup (op =) lits |> Option.map S |> the_default (V x)
- | resolve_sign_atom _ a = a
-fun resolve_mtype lits =
+fun resolve_annotation_atom asgs (V x) =
+ x |> AList.lookup (op =) asgs |> Option.map A |> the_default (V x)
+ | resolve_annotation_atom _ aa = aa
+fun resolve_mtype asgs =
let
fun aux MAlpha = MAlpha
- | aux (MFun (M1, a, M2)) = MFun (aux M1, resolve_sign_atom lits a, aux M2)
+ | aux (MFun (M1, aa, M2)) =
+ MFun (aux M1, resolve_annotation_atom asgs aa, aux M2)
| aux (MPair Mp) = MPair (pairself aux Mp)
| aux (MType (s, Ms)) = MType (s, map aux Ms)
| aux (MRec z) = MRec z
in aux end
-datatype comp_op = Eq | Leq
+datatype comp_op = Eq | Neq | Leq
-type comp = sign_atom * sign_atom * comp_op * var list
-type sign_expr = literal list
+type comp = annotation_atom * annotation_atom * comp_op * var list
+type assign_clause = assign_literal list
-type constraint_set = literal list * comp list * sign_expr list
+type constraint_set = comp list * assign_clause list
fun string_for_comp_op Eq = "="
+ | string_for_comp_op Neq = "\<noteq>"
| string_for_comp_op Leq = "\<le>"
-fun string_for_sign_expr [] = "\<bot>"
- | string_for_sign_expr lits =
- space_implode " \<or> " (map string_for_literal lits)
+fun string_for_comp (aa1, aa2, cmp, xs) =
+ string_for_annotation_atom aa1 ^ " " ^ string_for_comp_op cmp ^
+ subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_annotation_atom aa2
+
+fun string_for_assign_clause NONE = "\<top>"
+ | string_for_assign_clause (SOME []) = "\<bot>"
+ | string_for_assign_clause (SOME asgs) =
+ space_implode " \<or> " (map string_for_assign_literal asgs)
-fun do_literal _ NONE = NONE
- | do_literal (x, sn) (SOME lits) =
- case AList.lookup (op =) lits x of
- SOME sn' => if sn = sn' then SOME lits else NONE
- | NONE => SOME ((x, sn) :: lits)
+fun add_assign_literal (x, (sn, a)) clauses =
+ if exists (fn [(x', (sn', a'))] =>
+ x = x' andalso ((sn = sn' andalso a <> a') orelse
+ (sn <> sn' andalso a = a'))
+ | _ => false) clauses then
+ NONE
+ else
+ SOME ([(x, (sn, a))] :: clauses)
+
+fun add_assign_disjunct _ NONE = NONE
+ | add_assign_disjunct asg (SOME asgs) = SOME (insert (op =) asg asgs)
+
+fun add_assign_clause NONE = I
+ | add_assign_clause (SOME clause) = insert (op =) clause
-fun do_sign_atom_comp Eq [] a1 a2 (accum as (lits, comps)) =
- (case (a1, a2) of
- (S sn1, S sn2) => if sn1 = sn2 then SOME accum else NONE
- | (V x1, S sn2) =>
- Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits))
- | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps)
- | _ => do_sign_atom_comp Eq [] a2 a1 accum)
- | do_sign_atom_comp Leq [] a1 a2 (accum as (lits, comps)) =
- (case (a1, a2) of
- (_, S Minus) => SOME accum
- | (S Plus, _) => SOME accum
- | (S Minus, S Plus) => NONE
- | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps)
- | _ => do_sign_atom_comp Eq [] a1 a2 accum)
- | do_sign_atom_comp cmp xs a1 a2 (lits, comps) =
- SOME (lits, insert (op =) (a1, a2, cmp, xs) comps)
+fun annotation_comp Eq a1 a2 = (a1 = a2)
+ | annotation_comp Neq a1 a2 = (a1 <> a2)
+ | annotation_comp Leq a1 a2 = (a1 = a2 orelse a2 = Gen)
+
+fun sign_for_comp_op Eq = Plus
+ | sign_for_comp_op Neq = Minus
+ | sign_for_comp_op Leq = raise BAD ("sign_for_comp_op", "unexpected \"Leq\"")
+
+fun do_annotation_atom_comp Leq [] aa1 aa2 (cset as (comps, clauses)) =
+ (case (aa1, aa2) of
+ (A a1, A a2) => if annotation_comp Leq a1 a2 then SOME cset else NONE
+ | _ => SOME (insert (op =) (aa1, aa2, Leq, []) comps, clauses))
+ | do_annotation_atom_comp cmp [] aa1 aa2 (cset as (comps, clauses)) =
+ (case (aa1, aa2) of
+ (A a1, A a2) => if annotation_comp cmp a1 a2 then SOME cset else NONE
+ | (V x1, A a2) =>
+ clauses |> add_assign_literal (x1, (sign_for_comp_op cmp, a2))
+ |> Option.map (pair comps)
+ | (A _, V _) => do_annotation_atom_comp cmp [] aa2 aa1 cset
+ | (V _, V _) => SOME (insert (op =) (aa1, aa2, cmp, []) comps, clauses))
+ | do_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) =
+ SOME (insert (op =) (aa1, aa2, cmp, xs) comps, clauses)
+
+fun add_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) =
+ (trace_msg (fn () => "*** Add " ^ string_for_comp (aa1, aa2, cmp, xs));
+ case do_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) of
+ NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
+ | SOME cset => cset)
fun do_mtype_comp _ _ _ _ NONE = NONE
- | do_mtype_comp _ _ MAlpha MAlpha accum = accum
- | do_mtype_comp Eq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
- (SOME accum) =
- accum |> do_sign_atom_comp Eq xs a1 a2 |> do_mtype_comp Eq xs M11 M21
- |> do_mtype_comp Eq xs M12 M22
- | do_mtype_comp Leq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
- (SOME accum) =
+ | do_mtype_comp _ _ MAlpha MAlpha cset = cset
+ | do_mtype_comp Eq xs (MFun (M11, aa1, M12)) (MFun (M21, aa2, M22))
+ (SOME cset) =
+ cset |> do_annotation_atom_comp Eq xs aa1 aa2
+ |> do_mtype_comp Eq xs M11 M21 |> do_mtype_comp Eq xs M12 M22
+ | do_mtype_comp Leq xs (MFun (M11, aa1, M12)) (MFun (M21, aa2, M22))
+ (SOME cset) =
(if exists_alpha_sub_mtype M11 then
- accum |> do_sign_atom_comp Leq xs a1 a2
- |> do_mtype_comp Leq xs M21 M11
- |> (case a2 of
- S Minus => I
- | S Plus => do_mtype_comp Leq xs M11 M21
- | V x => do_mtype_comp Leq (x :: xs) M11 M21)
+ cset |> do_annotation_atom_comp Leq xs aa1 aa2
+ |> do_mtype_comp Leq xs M21 M11
+ |> (case aa2 of
+ A Gen => I
+ | A _ => do_mtype_comp Leq xs M11 M21
+ | V x => do_mtype_comp Leq (x :: xs) M11 M21)
else
- SOME accum)
+ SOME cset)
|> do_mtype_comp Leq xs M12 M22
| do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22))
- accum =
- (accum |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
+ cset =
+ (cset |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
handle ListPair.UnequalLengths =>
raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2], []))
- | do_mtype_comp _ _ (MType _) (MType _) accum =
- accum (* no need to compare them thanks to the cache *)
+ | do_mtype_comp _ _ (MType _) (MType _) cset =
+ cset (* no need to compare them thanks to the cache *)
| do_mtype_comp cmp _ M1 M2 _ =
raise MTYPE ("Nitpick_Mono.do_mtype_comp (" ^ string_for_comp_op cmp ^ ")",
[M1, M2], [])
-fun add_mtype_comp cmp M1 M2 ((lits, comps, sexps) : constraint_set) =
- (trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
- string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
- case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of
- NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
- | SOME (lits, comps) => (lits, comps, sexps))
+fun add_mtype_comp cmp M1 M2 cset =
+ (trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
+ string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
+ case SOME cset |> do_mtype_comp cmp [] M1 M2 of
+ NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
+ | SOME cset => cset)
val add_mtypes_equal = add_mtype_comp Eq
val add_is_sub_mtype = add_mtype_comp Leq
fun do_notin_mtype_fv _ _ _ NONE = NONE
- | do_notin_mtype_fv Minus _ MAlpha accum = accum
+ | do_notin_mtype_fv Minus _ MAlpha cset = cset
| do_notin_mtype_fv Plus [] MAlpha _ = NONE
- | do_notin_mtype_fv Plus [(x, sn)] MAlpha (SOME (lits, sexps)) =
- SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps)
- | do_notin_mtype_fv Plus sexp MAlpha (SOME (lits, sexps)) =
- SOME (lits, insert (op =) sexp sexps)
- | do_notin_mtype_fv sn sexp (MFun (M1, S sn', M2)) accum =
- accum |> (if sn' = Plus andalso sn = Plus then
- do_notin_mtype_fv Plus sexp M1
- else
- I)
- |> (if sn' = Minus orelse sn = Plus then
- do_notin_mtype_fv Minus sexp M1
- else
- I)
- |> do_notin_mtype_fv sn sexp M2
- | do_notin_mtype_fv Plus sexp (MFun (M1, V x, M2)) accum =
- accum |> (case do_literal (x, Minus) (SOME sexp) of
- NONE => I
- | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
- |> do_notin_mtype_fv Minus sexp M1
- |> do_notin_mtype_fv Plus sexp M2
- | do_notin_mtype_fv Minus sexp (MFun (M1, V x, M2)) accum =
- accum |> (case do_literal (x, Plus) (SOME sexp) of
- NONE => I
- | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
- |> do_notin_mtype_fv Minus sexp M2
- | do_notin_mtype_fv sn sexp (MPair (M1, M2)) accum =
- accum |> fold (do_notin_mtype_fv sn sexp) [M1, M2]
- | do_notin_mtype_fv sn sexp (MType (_, Ms)) accum =
- accum |> fold (do_notin_mtype_fv sn sexp) Ms
- | do_notin_mtype_fv _ _ M _ =
- raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
+ | do_notin_mtype_fv Plus [asg] MAlpha (SOME clauses) =
+ clauses |> add_assign_literal asg
+ | do_notin_mtype_fv Plus unless MAlpha (SOME clauses) =
+ SOME (insert (op =) unless clauses)
+ | do_notin_mtype_fv sn unless (MFun (M1, A a, M2)) cset =
+ cset |> (if a <> Gen andalso sn = Plus then do_notin_mtype_fv Plus unless M1
+ else I)
+ |> (if a = Gen orelse sn = Plus then do_notin_mtype_fv Minus unless M1
+ else I)
+ |> do_notin_mtype_fv sn unless M2
+ | do_notin_mtype_fv Plus unless (MFun (M1, V x, M2)) cset =
+ cset |> (case add_assign_disjunct (x, (Plus, Gen)) (SOME unless) of
+ NONE => I
+ | SOME unless' => do_notin_mtype_fv Plus unless' M1)
+ |> do_notin_mtype_fv Minus unless M1
+ |> do_notin_mtype_fv Plus unless M2
+ | do_notin_mtype_fv Minus unless (MFun (M1, V x, M2)) cset =
+ cset |> (case fold (fn a => add_assign_disjunct (x, (Plus, a))) [Fls, Tru]
+ (SOME unless) of
+ NONE => I
+ | SOME unless' => do_notin_mtype_fv Plus unless' M1)
+ |> do_notin_mtype_fv Minus unless M2
+ | do_notin_mtype_fv sn unless (MPair (M1, M2)) cset =
+ cset |> fold (do_notin_mtype_fv sn unless) [M1, M2]
+ | do_notin_mtype_fv sn unless (MType (_, Ms)) cset =
+ cset |> fold (do_notin_mtype_fv sn unless) Ms
+ | do_notin_mtype_fv _ _ M _ =
+ raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
-fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) =
- (trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
- (case sn of Minus => "concrete" | Plus => "complete"));
- case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
- NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
- | SOME (lits, sexps) => (lits, comps, sexps))
+fun add_notin_mtype_fv sn unless M (comps, clauses) =
+ (trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
+ (case sn of Minus => "concrete" | Plus => "complete"));
+ case SOME clauses |> do_notin_mtype_fv sn unless M of
+ NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
+ | SOME clauses => (comps, clauses))
val add_mtype_is_concrete = add_notin_mtype_fv Minus
val add_mtype_is_complete = add_notin_mtype_fv Plus
-val bool_from_minus = true
-
-fun bool_from_sign Plus = not bool_from_minus
- | bool_from_sign Minus = bool_from_minus
-fun sign_from_bool b = if b = bool_from_minus then Minus else Plus
+val bool_table =
+ [(Gen, (false, false)),
+ (New, (false, true)),
+ (Fls, (true, false)),
+ (Tru, (true, true))]
-fun prop_for_literal (x, sn) =
- (not (bool_from_sign sn) ? PropLogic.Not) (PropLogic.BoolVar x)
-fun prop_for_sign_atom_eq (S sn', sn) =
- if sn = sn' then PropLogic.True else PropLogic.False
- | prop_for_sign_atom_eq (V x, sn) = prop_for_literal (x, sn)
-fun prop_for_sign_expr xs = PropLogic.exists (map prop_for_literal xs)
-fun prop_for_exists_eq xs sn =
- PropLogic.exists (map (fn x => prop_for_literal (x, sn)) xs)
-fun prop_for_comp (a1, a2, Eq, []) =
- PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []),
- prop_for_comp (a2, a1, Leq, []))
- | prop_for_comp (a1, a2, Leq, []) =
- PropLogic.SOr (prop_for_sign_atom_eq (a1, Plus),
- prop_for_sign_atom_eq (a2, Minus))
- | prop_for_comp (a1, a2, cmp, xs) =
- PropLogic.SOr (prop_for_exists_eq xs Minus, prop_for_comp (a1, a2, cmp, []))
+fun fst_var n = 2 * n
+fun snd_var n = 2 * n + 1
-fun literals_from_assignments max_var assigns lits =
- fold (fn x => fn accum =>
- if AList.defined (op =) lits x then
- accum
- else case assigns x of
- SOME b => (x, sign_from_bool b) :: accum
- | NONE => accum) (max_var downto 1) lits
+val bools_from_annotation = AList.lookup (op =) bool_table #> the
+val annotation_from_bools = AList.find (op =) bool_table #> the_single
-fun string_for_comp (a1, a2, cmp, xs) =
- string_for_sign_atom a1 ^ " " ^ string_for_comp_op cmp ^
- subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_sign_atom a2
-
-fun print_problem lits comps sexps =
- trace_msg (fn () => "*** Problem:\n" ^
- cat_lines (map string_for_literal lits @
- map string_for_comp comps @
- map string_for_sign_expr sexps))
+fun prop_for_bool b = if b then PL.True else PL.False
+fun prop_for_bool_var_equality (v1, v2) =
+ PL.SAnd (PL.SOr (PL.BoolVar v1, PL.SNot (PL.BoolVar v2)),
+ PL.SOr (PL.SNot (PL.BoolVar v1), PL.BoolVar v2))
+fun prop_for_assign (x, a) =
+ let val (b1, b2) = bools_from_annotation a in
+ PL.SAnd (PL.BoolVar (fst_var x) |> not b1 ? PL.SNot,
+ PL.BoolVar (snd_var x) |> not b2 ? PL.SNot)
+ end
+fun prop_for_assign_literal (x, (Plus, a)) = prop_for_assign (x, a)
+ | prop_for_assign_literal (x, (Minus, a)) = PL.SNot (prop_for_assign (x, a))
+fun prop_for_atom_assign (A a', a) = prop_for_bool (a = a')
+ | prop_for_atom_assign (V x, a) = prop_for_assign_literal (x, (Plus, a))
+fun prop_for_atom_equality (aa1, A a2) = prop_for_atom_assign (aa1, a2)
+ | prop_for_atom_equality (A a1, aa2) = prop_for_atom_assign (aa2, a1)
+ | prop_for_atom_equality (V x1, V x2) =
+ PL.SAnd (prop_for_bool_var_equality (pairself fst_var (x1, x2)),
+ prop_for_bool_var_equality (pairself snd_var (x1, x2)))
+val prop_for_assign_clause = PL.exists o map prop_for_assign_literal
+fun prop_for_exists_var_assign_literal xs a =
+ PL.exists (map (fn x => prop_for_assign_literal (x, (Plus, a))) xs)
+fun prop_for_comp (aa1, aa2, Eq, []) =
+ PL.SAnd (prop_for_comp (aa1, aa2, Leq, []),
+ prop_for_comp (aa2, aa1, Leq, []))
+ | prop_for_comp (aa1, aa2, Neq, []) =
+ PL.SNot (prop_for_comp (aa1, aa2, Eq, []))
+ | prop_for_comp (aa1, aa2, Leq, []) =
+ PL.SOr (prop_for_atom_equality (aa1, aa2), prop_for_atom_assign (aa2, Gen))
+ | prop_for_comp (aa1, aa2, cmp, xs) =
+ PL.SOr (prop_for_exists_var_assign_literal xs Gen,
+ prop_for_comp (aa1, aa2, cmp, []))
-fun print_solution lits =
- let val (pos, neg) = List.partition (curry (op =) Plus o snd) lits in
- trace_msg (fn () => "*** Solution:\n" ^
- "+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^
- "-: " ^ commas (map (string_for_var o fst) neg))
- end
+fun extract_assigns max_var assigns asgs =
+ fold (fn x => fn accum =>
+ if AList.defined (op =) asgs x then
+ accum
+ else case (fst_var x, snd_var x) |> pairself assigns of
+ (NONE, NONE) => accum
+ | bp => (x, annotation_from_bools (pairself (the_default false) bp))
+ :: accum)
+ (max_var downto 1) asgs
+
+fun print_problem comps clauses =
+ trace_msg (fn () => "*** Problem:\n" ^
+ cat_lines (map string_for_comp comps @
+ map (string_for_assign_clause o SOME) clauses))
-fun solve max_var (lits, comps, sexps) =
+fun print_solution asgs =
+ trace_msg (fn () => "*** Solution:\n" ^
+ (asgs
+ |> map swap
+ |> AList.group (op =)
+ |> map (fn (a, xs) => string_for_annotation a ^ ": " ^
+ string_for_vars ", " (sort int_ord xs))
+ |> space_implode "\n"))
+
+(* The ML solver timeout should correspond more or less to the overhead of
+ invoking an external prover. *)
+val ml_solver_timeout = SOME (seconds 0.02)
+
+fun solve tac_timeout max_var (comps, clauses) =
let
+ val asgs =
+ map_filter (fn [(x, (Plus, a))] => SOME (x, a) | _ => NONE) clauses
fun do_assigns assigns =
- SOME (literals_from_assignments max_var assigns lits
- |> tap print_solution)
- val _ = print_problem lits comps sexps
- val prop = PropLogic.all (map prop_for_literal lits @
- map prop_for_comp comps @
- map prop_for_sign_expr sexps)
- val default_val = bool_from_sign Minus
+ SOME (extract_assigns max_var assigns asgs |> tap print_solution)
+ val _ = print_problem comps clauses
+ val prop =
+ PL.all (map prop_for_comp comps @ map prop_for_assign_clause clauses)
in
- if PropLogic.eval (K default_val) prop then
- do_assigns (K (SOME default_val))
+ if PL.eval (K false) prop then
+ do_assigns (K (SOME false))
+ else if PL.eval (K true) prop then
+ do_assigns (K (SOME true))
else
let
(* use the first ML solver (to avoid startup overhead) *)
- val solvers = !SatSolver.solvers
- |> filter (member (op =) ["dptsat", "dpll"] o fst)
+ val (ml_solvers, nonml_solvers) =
+ !SatSolver.solvers
+ |> List.partition (member (op =) ["dptsat", "dpll"] o fst)
+ val res =
+ if null nonml_solvers then
+ time_limit tac_timeout (snd (hd ml_solvers)) prop
+ else
+ time_limit ml_solver_timeout (snd (hd ml_solvers)) prop
+ handle TimeLimit.TimeOut =>
+ time_limit tac_timeout (SatSolver.invoke_solver "auto") prop
in
- case snd (hd solvers) prop of
+ case res of
SatSolver.SATISFIABLE assigns => do_assigns assigns
- | _ => NONE
+ | _ => (trace_msg (K "*** Unsolvable"); NONE)
end
+ handle TimeLimit.TimeOut => (trace_msg (K "*** Timed out"); NONE)
end
-type mtype_schema = mtyp * constraint_set
-type mtype_context =
+type mcontext =
{bound_Ts: typ list,
bound_Ms: mtyp list,
+ frame: (int * annotation_atom) list,
frees: (styp * mtyp) list,
consts: (styp * mtyp) list}
-type accumulator = mtype_context * constraint_set
+fun string_for_bound ctxt Ms (j, aa) =
+ Syntax.string_of_term ctxt (Bound (length Ms - j - 1)) ^ " :\<^bsup>" ^
+ string_for_annotation_atom aa ^ "\<^esup> " ^
+ string_for_mtype (nth Ms (length Ms - j - 1))
+fun string_for_free relevant_frees ((s, _), M) =
+ if member (op =) relevant_frees s then SOME (s ^ " : " ^ string_for_mtype M)
+ else NONE
+fun string_for_mcontext ctxt t {bound_Ms, frame, frees, ...} =
+ (map_filter (string_for_free (Term.add_free_names t [])) frees @
+ map (string_for_bound ctxt bound_Ms) frame)
+ |> commas |> enclose "[" "]"
+
+val initial_gamma =
+ {bound_Ts = [], bound_Ms = [], frame = [], frees = [], consts = []}
+
+fun push_bound aa T M {bound_Ts, bound_Ms, frame, frees, consts} =
+ {bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms,
+ frame = frame @ [(length bound_Ts, aa)], frees = frees, consts = consts}
+fun pop_bound {bound_Ts, bound_Ms, frame, frees, consts} =
+ {bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms,
+ frame = frame |> filter_out (fn (j, _) => j = length bound_Ts - 1),
+ frees = frees, consts = consts}
+ handle List.Empty => initial_gamma (* FIXME: needed? *)
+
+fun set_frame frame ({bound_Ts, bound_Ms, frees, consts, ...} : mcontext) =
+ {bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame, frees = frees,
+ consts = consts}
+
+(* FIXME: make sure tracing messages are complete *)
+
+fun add_comp_frame aa cmp = fold (add_annotation_atom_comp cmp [] aa o snd)
-val initial_gamma = {bound_Ts = [], bound_Ms = [], frees = [], consts = []}
+fun add_bound_frame j frame =
+ let
+ val (new_frame, gen_frame) = List.partition (curry (op =) j o fst) frame
+ in
+ add_comp_frame (A New) Leq new_frame
+ #> add_comp_frame (A Gen) Eq gen_frame
+ end
+
+fun fresh_frame ({max_fresh, ...} : mdata) fls tru =
+ map (apsnd (fn aa =>
+ case (aa, fls, tru) of
+ (A Fls, SOME a, _) => A a
+ | (A Tru, _, SOME a) => A a
+ | (A Gen, _, _) => A Gen
+ | _ => V (Unsynchronized.inc max_fresh)))
+
+fun conj_clauses res_aa aa1 aa2 =
+ [[(aa1, (Neq, Tru)), (aa2, (Neq, Tru)), (res_aa, (Eq, Tru))],
+ [(aa1, (Neq, Fls)), (res_aa, (Eq, Fls))],
+ [(aa2, (Neq, Fls)), (res_aa, (Eq, Fls))],
+ [(aa1, (Neq, Gen)), (aa2, (Eq, Fls)), (res_aa, (Eq, Gen))],
+ [(aa1, (Neq, New)), (aa2, (Eq, Fls)), (res_aa, (Eq, Gen))],
+ [(aa1, (Eq, Fls)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))],
+ [(aa1, (Eq, Fls)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]]
+
+fun disj_clauses res_aa aa1 aa2 =
+ [[(aa1, (Neq, Tru)), (res_aa, (Eq, Tru))],
+ [(aa2, (Neq, Tru)), (res_aa, (Eq, Tru))],
+ [(aa1, (Neq, Fls)), (aa2, (Neq, Fls)), (res_aa, (Eq, Fls))],
+ [(aa1, (Neq, Gen)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+ [(aa1, (Neq, New)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+ [(aa1, (Eq, Tru)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))],
+ [(aa1, (Eq, Tru)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]]
+
+fun imp_clauses res_aa aa1 aa2 =
+ [[(aa1, (Neq, Fls)), (res_aa, (Eq, Tru))],
+ [(aa2, (Neq, Tru)), (res_aa, (Eq, Tru))],
+ [(aa1, (Neq, Tru)), (aa2, (Neq, Fls)), (res_aa, (Eq, Fls))],
+ [(aa1, (Neq, Gen)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+ [(aa1, (Neq, New)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+ [(aa1, (Eq, Fls)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))],
+ [(aa1, (Eq, Fls)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]]
-fun push_bound T M {bound_Ts, bound_Ms, frees, consts} =
- {bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms, frees = frees,
- consts = consts}
-fun pop_bound {bound_Ts, bound_Ms, frees, consts} =
- {bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms, frees = frees,
- consts = consts}
- handle List.Empty => initial_gamma (* FIXME: needed? *)
+val meta_conj_triple = ("\<and>", conj_clauses, @{const Pure.conjunction})
+val meta_imp_triple = ("\<implies>", imp_clauses, @{const "==>"})
+val conj_triple = ("\<and>", conj_clauses, @{const conj})
+val disj_triple = ("\<or>", disj_clauses, @{const disj})
+val imp_triple = ("\<implies>", imp_clauses, @{const implies})
+
+fun add_annotation_clause_from_quasi_clause _ NONE = NONE
+ | add_annotation_clause_from_quasi_clause [] accum = accum
+ | add_annotation_clause_from_quasi_clause ((aa, (cmp, a)) :: rest) accum =
+ case aa of
+ A a' => if annotation_comp cmp a' a then NONE
+ else add_annotation_clause_from_quasi_clause rest accum
+ | V x => add_annotation_clause_from_quasi_clause rest accum
+ |> Option.map (cons (x, (sign_for_comp_op cmp, a)))
+
+fun assign_clause_from_quasi_clause unless =
+ add_annotation_clause_from_quasi_clause unless (SOME [])
+
+fun add_connective_var conn mk_quasi_clauses res_aa aa1 aa2 =
+ (trace_msg (fn () => "*** Add " ^ string_for_annotation_atom res_aa ^ " = " ^
+ string_for_annotation_atom aa1 ^ " " ^ conn ^ " " ^
+ string_for_annotation_atom aa2);
+ fold (add_assign_clause o assign_clause_from_quasi_clause)
+ (mk_quasi_clauses res_aa aa1 aa2))
+fun add_connective_frames conn mk_quasi_clauses res_frame frame1 frame2 =
+ fold I (map3 (fn (_, res_aa) => fn (_, aa1) => fn (_, aa2) =>
+ add_connective_var conn mk_quasi_clauses res_aa aa1 aa2)
+ res_frame frame1 frame2)
+
+fun kill_unused_in_frame is_in (accum as ({frame, ...}, _)) =
+ let val (used_frame, unused_frame) = List.partition is_in frame in
+ accum |>> set_frame used_frame
+ ||> add_comp_frame (A Gen) Eq unused_frame
+ end
+
+fun split_frame is_in_fun (gamma as {frame, ...}, cset) =
+ let
+ fun bubble fun_frame arg_frame [] cset =
+ ((rev fun_frame, rev arg_frame), cset)
+ | bubble fun_frame arg_frame ((bound as (_, aa)) :: rest) cset =
+ if is_in_fun bound then
+ bubble (bound :: fun_frame) arg_frame rest
+ (cset |> add_comp_frame aa Leq arg_frame)
+ else
+ bubble fun_frame (bound :: arg_frame) rest cset
+ in cset |> bubble [] [] frame ||> pair gamma end
+
+fun add_annotation_atom_comp_alt _ (A Gen) _ _ = I
+ | add_annotation_atom_comp_alt _ (A _) _ _ =
+ (trace_msg (K "*** Expected G"); raise UNSOLVABLE ())
+ | add_annotation_atom_comp_alt cmp (V x) aa1 aa2 =
+ add_annotation_atom_comp cmp [x] aa1 aa2
+
+fun add_arg_order1 ((_, aa), (_, prev_aa)) = I
+ add_annotation_atom_comp_alt Neq prev_aa (A Gen) aa
+fun add_app1 fun_aa ((_, res_aa), (_, arg_aa)) = I
+ let
+ val clause = [(arg_aa, (Eq, New)), (res_aa, (Eq, Gen))]
+ |> assign_clause_from_quasi_clause
+ in
+ trace_msg (fn () => "*** Add " ^ string_for_assign_clause clause);
+ apsnd (add_assign_clause clause)
+ #> add_annotation_atom_comp_alt Leq arg_aa fun_aa res_aa
+ end
+fun add_app _ [] [] = I
+ | add_app fun_aa res_frame arg_frame =
+ add_comp_frame (A New) Leq arg_frame
+ #> fold add_arg_order1 (tl arg_frame ~~ (fst (split_last arg_frame)))
+ #> fold (add_app1 fun_aa) (res_frame ~~ arg_frame)
+
+fun consider_connective mdata (conn, mk_quasi_clauses, t0) do_t1 do_t2
+ (accum as ({frame, ...}, _)) =
+ let
+ val mtype_for = fresh_mtype_for_type mdata false
+ val frame1 = fresh_frame mdata (SOME Tru) NONE frame
+ val frame2 = fresh_frame mdata (SOME Fls) NONE frame
+ val (m1, accum) = accum |>> set_frame frame1 |> do_t1
+ val (m2, accum) = accum |>> set_frame frame2 |> do_t2
+ in
+ (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
+ accum |>> set_frame frame
+ ||> apsnd (add_connective_frames conn mk_quasi_clauses frame frame1
+ frame2))
+ end
fun consider_term (mdata as {hol_ctxt = {thy, ctxt, stds, ...}, alpha_T,
max_fresh, ...}) =
@@ -557,21 +788,23 @@
<= length ts
| _ => true
val mtype_for = fresh_mtype_for_type mdata false
- fun plus_set_mtype_for_dom M =
- MFun (M, S (if exists_alpha_sub_mtype M then Plus else Minus), bool_M)
fun do_all T (gamma, cset) =
let
val abs_M = mtype_for (domain_type (domain_type T))
+ val x = Unsynchronized.inc max_fresh
val body_M = mtype_for (body_type T)
in
- (MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M),
- (gamma, cset |> add_mtype_is_complete abs_M))
+ (MFun (MFun (abs_M, V x, body_M), A Gen, body_M),
+ (gamma, cset |> add_mtype_is_complete [(x, (Plus, Tru))] abs_M))
end
fun do_equals T (gamma, cset) =
- let val M = mtype_for (domain_type T) in
- (MFun (M, S Minus, MFun (M, V (Unsynchronized.inc max_fresh),
- mtype_for (nth_range_type 2 T))),
- (gamma, cset |> add_mtype_is_concrete M))
+ let
+ val M = mtype_for (domain_type T)
+ val x = Unsynchronized.inc max_fresh
+ in
+ (MFun (M, A Gen, MFun (M, V x, mtype_for (nth_range_type 2 T))),
+ (gamma, cset |> add_mtype_is_concrete [] M
+ |> add_annotation_atom_comp Leq [] (A Fls) (V x)))
end
fun do_robust_set_operation T (gamma, cset) =
let
@@ -580,7 +813,7 @@
val M2 = mtype_for set_T
val M3 = mtype_for set_T
in
- (MFun (M1, S Minus, MFun (M2, S Minus, M3)),
+ (MFun (M1, A Gen, MFun (M2, A Gen, M3)),
(gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3))
end
fun do_fragile_set_operation T (gamma, cset) =
@@ -589,164 +822,175 @@
val set_M = mtype_for set_T
fun custom_mtype_for (T as Type (@{type_name fun}, [T1, T2])) =
if T = set_T then set_M
- else MFun (custom_mtype_for T1, S Minus, custom_mtype_for T2)
+ else MFun (custom_mtype_for T1, A Gen, custom_mtype_for T2)
| custom_mtype_for T = mtype_for T
in
- (custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete set_M))
+ (custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete [] set_M))
end
fun do_pair_constr T accum =
case mtype_for (nth_range_type 2 T) of
M as MPair (a_M, b_M) =>
- (MFun (a_M, S Minus, MFun (b_M, S Minus, M)), accum)
+ (MFun (a_M, A Gen, MFun (b_M, A Gen, M)), accum)
| M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M], [])
fun do_nth_pair_sel n T =
case mtype_for (domain_type T) of
M as MPair (a_M, b_M) =>
- pair (MFun (M, S Minus, if n = 0 then a_M else b_M))
+ pair (MFun (M, A Gen, if n = 0 then a_M else b_M))
| M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M], [])
- fun do_bounded_quantifier t0 abs_s abs_T connective_t bound_t body_t accum =
- let
- val abs_M = mtype_for abs_T
- val (bound_m, accum) =
- accum |>> push_bound abs_T abs_M |> do_term bound_t
- val expected_bound_M = plus_set_mtype_for_dom abs_M
- val (body_m, accum) =
- accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m)
- |> do_term body_t ||> apfst pop_bound
- val bound_M = mtype_of_mterm bound_m
- val (M1, a, _) = dest_MFun bound_M
- in
- (MApp (MRaw (t0, MFun (bound_M, S Minus, bool_M)),
- MAbs (abs_s, abs_T, M1, a,
- MApp (MApp (MRaw (connective_t,
- mtype_for (fastype_of connective_t)),
- MApp (bound_m, MRaw (Bound 0, M1))),
- body_m))), accum)
- end
- and do_term t (accum as ({bound_Ts, bound_Ms, frees, consts}, cset)) =
- (trace_msg (fn () => " \<Gamma> \<turnstile> " ^
- Syntax.string_of_term ctxt t ^ " : _?");
- case t of
- Const (x as (s, T)) =>
- (case AList.lookup (op =) consts x of
- SOME M => (M, accum)
- | NONE =>
- if not (could_exist_alpha_subtype alpha_T T) then
- (mtype_for T, accum)
- else case s of
- @{const_name all} => do_all T accum
- | @{const_name "=="} => do_equals T accum
- | @{const_name All} => do_all T accum
- | @{const_name Ex} =>
- let val set_T = domain_type T in
- do_term (Abs (Name.uu, set_T,
- @{const Not} $ (HOLogic.mk_eq
- (Abs (Name.uu, domain_type set_T,
- @{const False}),
- Bound 0)))) accum
- |>> mtype_of_mterm
- end
- | @{const_name HOL.eq} => do_equals T accum
- | @{const_name The} =>
- (trace_msg (K "*** The"); raise UNSOLVABLE ())
- | @{const_name Eps} =>
- (trace_msg (K "*** Eps"); raise UNSOLVABLE ())
- | @{const_name If} =>
- do_robust_set_operation (range_type T) accum
- |>> curry3 MFun bool_M (S Minus)
- | @{const_name Pair} => do_pair_constr T accum
- | @{const_name fst} => do_nth_pair_sel 0 T accum
- | @{const_name snd} => do_nth_pair_sel 1 T accum
- | @{const_name Id} =>
- (MFun (mtype_for (domain_type T), S Minus, bool_M), accum)
- | @{const_name converse} =>
+ and do_connect triple t1 t2 =
+ consider_connective mdata triple (do_term t1) (do_term t2)
+ and do_term t
+ (accum as (gamma as {bound_Ts, bound_Ms, frame, frees, consts},
+ cset)) =
+ (trace_msg (fn () => " " ^ string_for_mcontext ctxt t gamma ^
+ " \<turnstile> " ^ Syntax.string_of_term ctxt t ^
+ " : _?");
+ case t of
+ @{const False} =>
+ (MRaw (t, bool_M), accum ||> add_comp_frame (A Fls) Leq frame)
+ | Const (@{const_name None}, T) =>
+ (MRaw (t, mtype_for T), accum ||> add_comp_frame (A Fls) Leq frame)
+ | @{const True} =>
+ (MRaw (t, bool_M), accum ||> add_comp_frame (A Tru) Leq frame)
+ | (t0 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t2 =>
+ (* hack to exploit symmetry of equality when typing "insert" *)
+ (if t2 = Bound 0 then do_term @{term True}
+ else do_term (t0 $ t2 $ Bound 0)) accum
+ | Const (x as (s, T)) =>
+ (case AList.lookup (op =) consts x of
+ SOME M => (M, accum)
+ | NONE =>
+ if not (could_exist_alpha_subtype alpha_T T) then
+ (mtype_for T, accum)
+ else case s of
+ @{const_name all} => do_all T accum
+ | @{const_name "=="} => do_equals T accum
+ | @{const_name All} => do_all T accum
+ | @{const_name Ex} =>
+ let val set_T = domain_type T in
+ do_term (Abs (Name.uu, set_T,
+ @{const Not} $ (HOLogic.mk_eq
+ (Abs (Name.uu, domain_type set_T,
+ @{const False}),
+ Bound 0)))) accum
+ |>> mtype_of_mterm
+ end
+ | @{const_name HOL.eq} => do_equals T accum
+ | @{const_name The} =>
+ (trace_msg (K "*** The"); raise UNSOLVABLE ())
+ | @{const_name Eps} =>
+ (trace_msg (K "*** Eps"); raise UNSOLVABLE ())
+ | @{const_name If} =>
+ do_robust_set_operation (range_type T) accum
+ |>> curry3 MFun bool_M (A Gen)
+ | @{const_name Pair} => do_pair_constr T accum
+ | @{const_name fst} => do_nth_pair_sel 0 T accum
+ | @{const_name snd} => do_nth_pair_sel 1 T accum
+ | @{const_name Id} =>
+ (MFun (mtype_for (domain_type T), A Gen, bool_M), accum)
+ | @{const_name converse} =>
+ let
+ val x = Unsynchronized.inc max_fresh
+ fun mtype_for_set T =
+ MFun (mtype_for (domain_type T), V x, bool_M)
+ val ab_set_M = domain_type T |> mtype_for_set
+ val ba_set_M = range_type T |> mtype_for_set
+ in
+ (MFun (ab_set_M, A Gen, ba_set_M),
+ accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+ end
+ | @{const_name trancl} => do_fragile_set_operation T accum
+ | @{const_name rel_comp} =>
+ let
+ val x = Unsynchronized.inc max_fresh
+ fun mtype_for_set T =
+ MFun (mtype_for (domain_type T), V x, bool_M)
+ val bc_set_M = domain_type T |> mtype_for_set
+ val ab_set_M = domain_type (range_type T) |> mtype_for_set
+ val ac_set_M = nth_range_type 2 T |> mtype_for_set
+ in
+ (MFun (bc_set_M, A Gen, MFun (ab_set_M, A Gen, ac_set_M)),
+ accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+ end
+ | @{const_name image} =>
+ let
+ val x = Unsynchronized.inc max_fresh
+ val a_M = mtype_for (domain_type (domain_type T))
+ val b_M = mtype_for (range_type (domain_type T))
+ in
+ (MFun (MFun (a_M, A Gen, b_M), A Gen,
+ MFun (MFun (a_M, V x, bool_M), A Gen,
+ MFun (b_M, V x, bool_M))),
+ accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+ end
+ | @{const_name finite} =>
+ let
+ val M1 = mtype_for (domain_type (domain_type T))
+ val a = if exists_alpha_sub_mtype M1 then Fls else Gen
+ in (MFun (MFun (M1, A a, bool_M), A Gen, bool_M), accum) end
+ | @{const_name Sigma} =>
+ let
+ val x = Unsynchronized.inc max_fresh
+ fun mtype_for_set T =
+ MFun (mtype_for (domain_type T), V x, bool_M)
+ val a_set_T = domain_type T
+ val a_M = mtype_for (domain_type a_set_T)
+ val b_set_M =
+ mtype_for_set (range_type (domain_type (range_type T)))
+ val a_set_M = mtype_for_set a_set_T
+ val a_to_b_set_M = MFun (a_M, A Gen, b_set_M)
+ val ab_set_M = mtype_for_set (nth_range_type 2 T)
+ in
+ (MFun (a_set_M, A Gen, MFun (a_to_b_set_M, A Gen, ab_set_M)),
+ accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+ end
+ | _ =>
+ if s = @{const_name safe_The} then
let
- val x = Unsynchronized.inc max_fresh
- fun mtype_for_set T =
- MFun (mtype_for (domain_type T), V x, bool_M)
- val ab_set_M = domain_type T |> mtype_for_set
- val ba_set_M = range_type T |> mtype_for_set
- in (MFun (ab_set_M, S Minus, ba_set_M), accum) end
- | @{const_name trancl} => do_fragile_set_operation T accum
- | @{const_name rel_comp} =>
- let
- val x = Unsynchronized.inc max_fresh
- fun mtype_for_set T =
- MFun (mtype_for (domain_type T), V x, bool_M)
- val bc_set_M = domain_type T |> mtype_for_set
- val ab_set_M = domain_type (range_type T) |> mtype_for_set
- val ac_set_M = nth_range_type 2 T |> mtype_for_set
- in
- (MFun (bc_set_M, S Minus, MFun (ab_set_M, S Minus, ac_set_M)),
- accum)
- end
- | @{const_name image} =>
- let
- val a_M = mtype_for (domain_type (domain_type T))
- val b_M = mtype_for (range_type (domain_type T))
- in
- (MFun (MFun (a_M, S Minus, b_M), S Minus,
- MFun (plus_set_mtype_for_dom a_M, S Minus,
- plus_set_mtype_for_dom b_M)), accum)
- end
- | @{const_name finite} =>
- let val M1 = mtype_for (domain_type (domain_type T)) in
- (MFun (plus_set_mtype_for_dom M1, S Minus, bool_M), accum)
- end
- | @{const_name Sigma} =>
- let
- val x = Unsynchronized.inc max_fresh
- fun mtype_for_set T =
- MFun (mtype_for (domain_type T), V x, bool_M)
- val a_set_T = domain_type T
- val a_M = mtype_for (domain_type a_set_T)
- val b_set_M = mtype_for_set (range_type (domain_type
- (range_type T)))
- val a_set_M = mtype_for_set a_set_T
- val a_to_b_set_M = MFun (a_M, S Minus, b_set_M)
- val ab_set_M = mtype_for_set (nth_range_type 2 T)
- in
- (MFun (a_set_M, S Minus,
- MFun (a_to_b_set_M, S Minus, ab_set_M)), accum)
- end
- | _ =>
- if s = @{const_name safe_The} then
- let
- val a_set_M = mtype_for (domain_type T)
- val a_M = dest_MFun a_set_M |> #1
- in (MFun (a_set_M, S Minus, a_M), accum) end
- else if s = @{const_name ord_class.less_eq} andalso
- is_set_type (domain_type T) then
- do_fragile_set_operation T accum
- else if is_sel s then
- (mtype_for_sel mdata x, accum)
- else if is_constr ctxt stds x then
- (mtype_for_constr mdata x, accum)
- else if is_built_in_const thy stds x then
- (fresh_mtype_for_type mdata true T, accum)
- else
- let val M = mtype_for T in
- (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
- frees = frees, consts = (x, M) :: consts}, cset))
- end) |>> curry MRaw t
+ val a_set_M = mtype_for (domain_type T)
+ val a_M = dest_MFun a_set_M |> #1
+ in (MFun (a_set_M, A Gen, a_M), accum) end
+ else if s = @{const_name ord_class.less_eq} andalso
+ is_set_type (domain_type T) then
+ do_fragile_set_operation T accum
+ else if is_sel s then
+ (mtype_for_sel mdata x, accum)
+ else if is_constr ctxt stds x then
+ (mtype_for_constr mdata x, accum)
+ else if is_built_in_const thy stds x then
+ (fresh_mtype_for_type mdata true T, accum)
+ else
+ let val M = mtype_for T in
+ (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame,
+ frees = frees, consts = (x, M) :: consts}, cset))
+ end)
+ |>> curry MRaw t
+ ||> apsnd (add_comp_frame (A Gen) Eq frame)
| Free (x as (_, T)) =>
(case AList.lookup (op =) frees x of
SOME M => (M, accum)
| NONE =>
let val M = mtype_for T in
- (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
+ (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame,
frees = (x, M) :: frees, consts = consts}, cset))
- end) |>> curry MRaw t
+ end)
+ |>> curry MRaw t ||> apsnd (add_comp_frame (A Gen) Eq frame)
| Var _ => (trace_msg (K "*** Var"); raise UNSOLVABLE ())
- | Bound j => (MRaw (t, nth bound_Ms j), accum)
+ | Bound j =>
+ (MRaw (t, nth bound_Ms j),
+ accum ||> add_bound_frame (length bound_Ts - j - 1) frame)
| Abs (s, T, t') =>
(case fin_fun_body T (fastype_of1 (T :: bound_Ts, t')) t' of
SOME t' =>
let
val M = mtype_for T
- val a = V (Unsynchronized.inc max_fresh)
- val (m', accum) = do_term t' (accum |>> push_bound T M)
- in (MAbs (s, T, M, a, m'), accum |>> pop_bound) end
+ val x = Unsynchronized.inc max_fresh
+ val (m', accum) = do_term t' (accum |>> push_bound (V x) T M)
+ in
+ (MAbs (s, T, M, V x, m'),
+ accum |>> pop_bound
+ ||> add_annotation_atom_comp Leq [] (A Fls) (V x))
+ end
| NONE =>
((case t' of
t1' $ Bound 0 =>
@@ -764,38 +1008,49 @@
handle SAME () =>
let
val M = mtype_for T
- val (m', accum) = do_term t' (accum |>> push_bound T M)
- in
- (MAbs (s, T, M, S Minus, m'), accum |>> pop_bound)
- end))
- | (t0 as Const (@{const_name All}, _))
- $ Abs (s', T', (t10 as @{const HOL.implies}) $ (t11 $ Bound 0) $ t12) =>
- do_bounded_quantifier t0 s' T' t10 t11 t12 accum
- | (t0 as Const (@{const_name Ex}, _))
- $ Abs (s', T', (t10 as @{const HOL.conj}) $ (t11 $ Bound 0) $ t12) =>
- do_bounded_quantifier t0 s' T' t10 t11 t12 accum
+ val x = Unsynchronized.inc max_fresh
+ val (m', accum) =
+ do_term t' (accum |>> push_bound (V x) T M)
+ in (MAbs (s, T, M, V x, m'), accum |>> pop_bound) end))
+ | @{const Not} $ t1 => do_connect imp_triple t1 @{const False} accum
+ | @{const conj} $ t1 $ t2 => do_connect conj_triple t1 t2 accum
+ | @{const disj} $ t1 $ t2 => do_connect disj_triple t1 t2 accum
+ | @{const implies} $ t1 $ t2 => do_connect imp_triple t1 t2 accum
| Const (@{const_name Let}, _) $ t1 $ t2 =>
do_term (betapply (t2, t1)) accum
| t1 $ t2 =>
let
- val (m1, accum) = do_term t1 accum
- val (m2, accum) = do_term t2 accum
+ fun is_in t (j, _) = loose_bvar1 (t, length bound_Ts - j - 1)
+ val accum as ({frame, ...}, _) =
+ accum |> kill_unused_in_frame (is_in t)
+ val ((frame1a, frame1b), accum) = accum |> split_frame (is_in t1)
+ val frame2a = frame1a |> map (apsnd (K (A Gen)))
+ val frame2b =
+ frame1b |> map (apsnd (fn _ => V (Unsynchronized.inc max_fresh)))
+ val frame2 = frame2a @ frame2b
+ val (m1, accum) = accum |>> set_frame frame1a |> do_term t1
+ val (m2, accum) = accum |>> set_frame frame2 |> do_term t2
in
let
- val M11 = mtype_of_mterm m1 |> dest_MFun |> #1
+ val (M11, aa, _) = mtype_of_mterm m1 |> dest_MFun
val M2 = mtype_of_mterm m2
- in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
+ in
+ (MApp (m1, m2),
+ accum |>> set_frame frame
+ ||> add_is_sub_mtype M2 M11
+ ||> add_app aa frame1b frame2b)
+ end
end)
- |> tap (fn (m, _) => trace_msg (fn () => " \<Gamma> \<turnstile> " ^
- string_for_mterm ctxt m))
+ |> tap (fn (m, (gamma, _)) =>
+ trace_msg (fn () => " " ^ string_for_mcontext ctxt t gamma ^
+ " \<turnstile> " ^
+ string_for_mterm ctxt m))
in do_term end
-fun force_minus_funs 0 _ = I
- | force_minus_funs n (M as MFun (M1, _, M2)) =
- add_mtypes_equal M (MFun (M1, S Minus, M2))
- #> force_minus_funs (n - 1) M2
- | force_minus_funs _ M =
- raise MTYPE ("Nitpick_Mono.force_minus_funs", [M], [])
+fun force_gen_funs 0 _ = I
+ | force_gen_funs n (M as MFun (M1, _, M2)) =
+ add_mtypes_equal M (MFun (M1, A Gen, M2)) #> force_gen_funs (n - 1) M2
+ | force_gen_funs _ M = raise MTYPE ("Nitpick_Mono.force_gen_funs", [M], [])
fun consider_general_equals mdata def (x as (_, T)) t1 t2 accum =
let
val (m1, accum) = consider_term mdata t1 accum
@@ -805,97 +1060,89 @@
val accum = accum ||> add_mtypes_equal M1 M2
val body_M = fresh_mtype_for_type mdata false (nth_range_type 2 T)
val m = MApp (MApp (MRaw (Const x,
- MFun (M1, S Minus, MFun (M2, S Minus, body_M))), m1), m2)
+ MFun (M1, A Gen, MFun (M2, A Gen, body_M))), m1), m2)
in
- (m, if def then
- let val (head_m, arg_ms) = strip_mcomb m1 in
- accum ||> force_minus_funs (length arg_ms) (mtype_of_mterm head_m)
- end
- else
- accum)
+ (m, (if def then
+ let val (head_m, arg_ms) = strip_mcomb m1 in
+ accum ||> force_gen_funs (length arg_ms) (mtype_of_mterm head_m)
+ end
+ else
+ accum))
end
-fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) =
+fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, max_fresh,
+ ...}) =
let
val mtype_for = fresh_mtype_for_type mdata false
val do_term = consider_term mdata
- fun do_formula sn t accum =
- let
- fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
- let
- val abs_M = mtype_for abs_T
- val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
- val (body_m, accum) =
- accum ||> side_cond ? add_mtype_is_complete abs_M
- |>> push_bound abs_T abs_M |> do_formula sn body_t
- val body_M = mtype_of_mterm body_m
- in
- (MApp (MRaw (Const quant_x,
- MFun (MFun (abs_M, S Minus, body_M), S Minus,
- body_M)),
- MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
- accum |>> pop_bound)
- end
- fun do_equals x t1 t2 =
- case sn of
- Plus => do_term t accum
- | Minus => consider_general_equals mdata false x t1 t2 accum
- in
- (trace_msg (fn () => " \<Gamma> \<turnstile> " ^
- Syntax.string_of_term ctxt t ^ " : o\<^sup>" ^
- string_for_sign sn ^ "?");
- case t of
- Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
- do_quantifier x s1 T1 t1
- | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
- | @{const Trueprop} $ t1 =>
- let val (m1, accum) = do_formula sn t1 accum in
- (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
- m1), accum)
- end
- | @{const Not} $ t1 =>
- let val (m1, accum) = do_formula (negate sn) t1 accum in
- (MApp (MRaw (@{const Not}, mtype_for (bool_T --> bool_T)), m1),
- accum)
- end
- | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
- do_quantifier x s1 T1 t1
- | Const (x0 as (@{const_name Ex}, T0))
- $ (t1 as Abs (s1, T1, t1')) =>
- (case sn of
- Plus => do_quantifier x0 s1 T1 t1'
- | Minus =>
- (* FIXME: Move elsewhere *)
- do_term (@{const Not}
- $ (HOLogic.eq_const (domain_type T0) $ t1
- $ Abs (Name.uu, T1, @{const False}))) accum)
- | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
- do_equals x t1 t2
- | Const (@{const_name Let}, _) $ t1 $ t2 =>
- do_formula sn (betapply (t2, t1)) accum
- | (t0 as Const (s0, _)) $ t1 $ t2 =>
- if s0 = @{const_name "==>"} orelse
- s0 = @{const_name Pure.conjunction} orelse
- s0 = @{const_name HOL.conj} orelse
- s0 = @{const_name HOL.disj} orelse
- s0 = @{const_name HOL.implies} then
- let
- val impl = (s0 = @{const_name "==>"} orelse
- s0 = @{const_name HOL.implies})
- val (m1, accum) = do_formula (sn |> impl ? negate) t1 accum
- val (m2, accum) = do_formula sn t2 accum
- in
- (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
- accum)
- end
- else
- do_term t accum
- | _ => do_term t accum)
- end
- |> tap (fn (m, _) =>
- trace_msg (fn () => "\<Gamma> \<turnstile> " ^
- string_for_mterm ctxt m ^ " : o\<^sup>" ^
- string_for_sign sn))
+ fun do_formula sn t (accum as (gamma, _)) =
+ let
+ fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
+ let
+ val abs_M = mtype_for abs_T
+ val x = Unsynchronized.inc max_fresh
+ val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
+ fun ann () = if quant_s = @{const_name Ex} then Fls else Tru
+ val (body_m, accum) =
+ accum ||> side_cond
+ ? add_mtype_is_complete [(x, (Plus, ann ()))] abs_M
+ |>> push_bound (V x) abs_T abs_M |> do_formula sn body_t
+ val body_M = mtype_of_mterm body_m
+ in
+ (MApp (MRaw (Const quant_x,
+ MFun (MFun (abs_M, A Gen, body_M), A Gen, body_M)),
+ MAbs (abs_s, abs_T, abs_M, A Gen, body_m)),
+ accum |>> pop_bound)
+ end
+ fun do_connect triple neg1 t1 t2 =
+ consider_connective mdata triple
+ (do_formula (sn |> neg1 ? negate_sign) t1) (do_formula sn t2)
+ fun do_equals x t1 t2 =
+ case sn of
+ Plus => do_term t accum
+ | Minus => consider_general_equals mdata false x t1 t2 accum
+ in
+ trace_msg (fn () => " " ^ string_for_mcontext ctxt t gamma ^
+ " \<turnstile> " ^ Syntax.string_of_term ctxt t ^
+ " : o\<^sup>" ^ string_for_sign sn ^ "?");
+ case t of
+ Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
+ do_quantifier x s1 T1 t1
+ | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
+ | @{const Trueprop} $ t1 =>
+ let val (m1, accum) = do_formula sn t1 accum in
+ (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)), m1),
+ accum)
+ end
+ | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
+ do_quantifier x s1 T1 t1
+ | Const (x0 as (@{const_name Ex}, T0)) $ (t1 as Abs (s1, T1, t1')) =>
+ (case sn of
+ Plus => do_quantifier x0 s1 T1 t1'
+ | Minus =>
+ (* FIXME: Move elsewhere *)
+ do_term (@{const Not}
+ $ (HOLogic.eq_const (domain_type T0) $ t1
+ $ Abs (Name.uu, T1, @{const False}))) accum)
+ | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 => do_equals x t1 t2
+ | Const (@{const_name Let}, _) $ t1 $ t2 =>
+ do_formula sn (betapply (t2, t1)) accum
+ | @{const Pure.conjunction} $ t1 $ t2 =>
+ do_connect meta_conj_triple false t1 t2 accum
+ | @{const "==>"} $ t1 $ t2 =>
+ do_connect meta_imp_triple true t1 t2 accum
+ | @{const Not} $ t1 =>
+ do_connect imp_triple true t1 @{const False} accum
+ | @{const conj} $ t1 $ t2 => do_connect conj_triple false t1 t2 accum
+ | @{const disj} $ t1 $ t2 => do_connect disj_triple false t1 t2 accum
+ | @{const implies} $ t1 $ t2 => do_connect imp_triple true t1 t2 accum
+ | _ => do_term t accum
+ end
+ |> tap (fn (m, (gamma, _)) =>
+ trace_msg (fn () => string_for_mcontext ctxt t gamma ^
+ " \<turnstile> " ^
+ string_for_mterm ctxt m ^ " : o\<^sup>" ^
+ string_for_sign sn))
in do_formula end
(* The harmless axiom optimization below is somewhat too aggressive in the face
@@ -928,12 +1175,12 @@
let
val abs_M = mtype_for abs_T
val (body_m, accum) =
- accum |>> push_bound abs_T abs_M |> do_formula body_t
+ accum |>> push_bound (A Gen) abs_T abs_M |> do_formula body_t
val body_M = mtype_of_mterm body_m
in
- (MApp (MRaw (quant_t,
- MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M)),
- MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
+ (MApp (MRaw (quant_t, MFun (MFun (abs_M, A Gen, body_M), A Gen,
+ body_M)),
+ MAbs (abs_s, abs_T, abs_M, A Gen, body_m)),
accum |>> pop_bound)
end
and do_conjunction t0 t1 t2 accum =
@@ -951,50 +1198,50 @@
(MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
end
and do_formula t accum =
- case t of
- (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
- do_all t0 s1 T1 t1 accum
- | @{const Trueprop} $ t1 =>
- let val (m1, accum) = do_formula t1 accum in
- (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
- m1), accum)
- end
- | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
- consider_general_equals mdata true x t1 t2 accum
- | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
- | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
- do_conjunction t0 t1 t2 accum
- | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
- do_all t0 s0 T1 t1 accum
- | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
- consider_general_equals mdata true x t1 t2 accum
- | (t0 as @{const HOL.conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
- | (t0 as @{const HOL.implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
- | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
- \do_formula", [t])
+ case t of
+ (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
+ do_all t0 s1 T1 t1 accum
+ | @{const Trueprop} $ t1 =>
+ let val (m1, accum) = do_formula t1 accum in
+ (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)), m1),
+ accum)
+ end
+ | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
+ consider_general_equals mdata true x t1 t2 accum
+ | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+ | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
+ do_conjunction t0 t1 t2 accum
+ | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
+ do_all t0 s0 T1 t1 accum
+ | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
+ consider_general_equals mdata true x t1 t2 accum
+ | (t0 as @{const conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
+ | (t0 as @{const implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+ | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
+ \do_formula", [t])
in do_formula t end
-fun string_for_mtype_of_term ctxt lits t M =
- Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype lits M)
+fun string_for_mtype_of_term ctxt asgs t M =
+ Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype asgs M)
-fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) =
+fun print_mcontext ctxt asgs ({frees, consts, ...} : mcontext) =
trace_msg (fn () =>
- map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
- map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
+ map (fn (x, M) => string_for_mtype_of_term ctxt asgs (Free x) M) frees @
+ map (fn (x, M) => string_for_mtype_of_term ctxt asgs (Const x) M) consts
|> cat_lines)
fun amass f t (ms, accum) =
let val (m, accum) = f t accum in (m :: ms, accum) end
-fun infer which no_harmless (hol_ctxt as {ctxt, ...}) binarize alpha_T
- (nondef_ts, def_ts) =
+fun infer which no_harmless (hol_ctxt as {ctxt, tac_timeout, ...}) binarize
+ alpha_T (nondef_ts, def_ts) =
let
val _ = trace_msg (fn () => "****** " ^ which ^ " analysis: " ^
string_for_mtype MAlpha ^ " is " ^
Syntax.string_of_typ ctxt alpha_T)
val mdata as {max_fresh, constr_mcache, ...} =
initial_mdata hol_ctxt binarize no_harmless alpha_T
- val accum = (initial_gamma, ([], [], []))
+ val accum = (initial_gamma, ([], []))
val (nondef_ms, accum) =
([], accum) |> amass (consider_general_formula mdata Plus) (hd nondef_ts)
|> fold (amass (consider_nondefinitional_axiom mdata))
@@ -1002,9 +1249,9 @@
val (def_ms, (gamma, cset)) =
([], accum) |> fold (amass (consider_definitional_axiom mdata)) def_ts
in
- case solve (!max_fresh) cset of
- SOME lits => (print_mtype_context ctxt lits gamma;
- SOME (lits, (nondef_ms, def_ms), !constr_mcache))
+ case solve tac_timeout (!max_fresh) cset of
+ SOME asgs => (print_mcontext ctxt asgs gamma;
+ SOME (asgs, (nondef_ms, def_ms), !constr_mcache))
| _ => NONE
end
handle UNSOLVABLE () => NONE
@@ -1019,23 +1266,23 @@
fun fin_fun_constr T1 T2 =
(@{const_name FinFun}, (T1 --> T2) --> Type (@{type_name fin_fun}, [T1, T2]))
-fun finitize_funs (hol_ctxt as {thy, ctxt, stds, constr_cache, ...})
- binarize finitizes alpha_T tsp =
+fun finitize_funs (hol_ctxt as {thy, ctxt, stds, constr_cache, ...}) binarize
+ finitizes alpha_T tsp =
case infer "Finiteness" true hol_ctxt binarize alpha_T tsp of
- SOME (lits, msp, constr_mtypes) =>
- if forall (curry (op =) Minus o snd) lits then
+ SOME (asgs, msp, constr_mtypes) =>
+ if forall (curry (op =) Gen o snd) asgs then
tsp
else
let
- fun should_finitize T a =
+ fun should_finitize T aa =
case triple_lookup (type_match thy) finitizes T of
SOME (SOME false) => false
- | _ => resolve_sign_atom lits a = S Plus
+ | _ => resolve_annotation_atom asgs aa = A Fls
fun type_from_mtype T M =
case (M, T) of
(MAlpha, _) => T
- | (MFun (M1, a, M2), Type (@{type_name fun}, Ts)) =>
- Type (if should_finitize T a then @{type_name fin_fun}
+ | (MFun (M1, aa, M2), Type (@{type_name fun}, Ts)) =>
+ Type (if should_finitize T aa then @{type_name fin_fun}
else @{type_name fun}, map2 type_from_mtype Ts [M1, M2])
| (MPair (M1, M2), Type (@{type_name prod}, Ts)) =>
Type (@{type_name prod}, map2 type_from_mtype Ts [M1, M2])
@@ -1106,14 +1353,14 @@
| _ => raise TYPE ("Nitpick_Mono.finitize_funs.term_from_mterm",
[T1], [])
in betapply (t1', coerce_term hol_ctxt new_Ts T2' T2 t2) end
- | MAbs (s, old_T, M, a, m') =>
+ | MAbs (s, old_T, M, aa, m') =>
let
val new_T = type_from_mtype old_T M
val t' = term_from_mterm (new_T :: new_Ts) (old_T :: old_Ts) m'
val T' = fastype_of1 (new_T :: new_Ts, t')
in
Abs (s, new_T, t')
- |> should_finitize (new_T --> T') a
+ |> should_finitize (new_T --> T') aa
? construct_value ctxt stds (fin_fun_constr new_T T') o single
end
in