File ‹refute.ML›
signature REFUTE =
sig
exception REFUTE of string * string
type params
type interpretation
type model
type arguments
exception MAXVARS_EXCEEDED
val add_interpreter : string -> (Proof.context -> model -> arguments -> term ->
(interpretation * model * arguments) option) -> theory -> theory
val add_printer : string -> (Proof.context -> model -> typ ->
interpretation -> (int -> bool) -> term option) -> theory -> theory
val interpret : Proof.context -> model -> arguments -> term ->
(interpretation * model * arguments)
val print : Proof.context -> model -> typ -> interpretation -> (int -> bool) -> term
val print_model : Proof.context -> model -> (int -> bool) -> string
val set_default_param : (string * string) -> theory -> theory
val get_default_param : Proof.context -> string -> string option
val get_default_params : Proof.context -> (string * string) list
val actual_params : Proof.context -> (string * string) list -> params
val find_model :
Proof.context -> params -> term list -> term -> bool -> string
val satisfy_term :
Proof.context -> (string * string) list -> term list -> term -> string
val refute_term :
Proof.context -> (string * string) list -> term list -> term -> string
val refute_goal :
Proof.context -> (string * string) list -> thm -> int -> string
val get_classdef : theory -> string -> (string * term) option
val norm_rhs : term -> term
val get_def : theory -> string * typ -> (string * term) option
val get_typedef : theory -> typ -> (string * term) option
val is_IDT_constructor : theory -> string * typ -> bool
val is_IDT_recursor : theory -> string * typ -> bool
val is_const_of_class: theory -> string * typ -> bool
val string_of_typ : typ -> string
end;
structure Refute : REFUTE =
struct
open Prop_Logic;
exception REFUTE of string * string;
exception MAXVARS_EXCEEDED;
datatype 'a tree =
Leaf of 'a
| Node of ('a tree) list;
fun tree_map f tr =
case tr of
Leaf x => Leaf (f x)
| Node xs => Node (map (tree_map f) xs);
fun tree_pair (t1, t2) =
case t1 of
Leaf x =>
(case t2 of
Leaf y => Leaf (x,y)
| Node _ => raise REFUTE ("tree_pair",
"trees are of different height (second tree is higher)"))
| Node xs =>
(case t2 of
Node ys => Node (map tree_pair (xs ~~ ys))
| Leaf _ => raise REFUTE ("tree_pair",
"trees are of different height (first tree is higher)"));
type params =
{
sizes : (string * int) list,
minsize : int,
maxsize : int,
maxvars : int,
maxtime : int,
satsolver: string,
no_assms : bool,
expect : string
};
type interpretation =
prop_formula list tree;
type model =
(typ * int) list * (term * interpretation) list;
type arguments =
{
maxvars : int,
def_eq : bool,
next_idx : int,
bounds : interpretation list,
wellformed: prop_formula
};
structure Data = Theory_Data
(
type T =
{interpreters: (string * (Proof.context -> model -> arguments -> term ->
(interpretation * model * arguments) option)) list,
printers: (string * (Proof.context -> model -> typ -> interpretation ->
(int -> bool) -> term option)) list,
parameters: string Symtab.table};
val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
fun merge
({interpreters = in1, printers = pr1, parameters = pa1},
{interpreters = in2, printers = pr2, parameters = pa2}) : T =
{interpreters = AList.merge (op =) (K true) (in1, in2),
printers = AList.merge (op =) (K true) (pr1, pr2),
parameters = Symtab.merge (op =) (pa1, pa2)};
);
val get_data = Data.get o Proof_Context.theory_of;
fun interpret ctxt model args t =
case get_first (fn (_, f) => f ctxt model args t)
(#interpreters (get_data ctxt)) of
NONE => raise REFUTE ("interpret",
"no interpreter for term " ^ quote (Syntax.string_of_term ctxt t))
| SOME x => x;
fun print ctxt model T intr assignment =
case get_first (fn (_, f) => f ctxt model T intr assignment)
(#printers (get_data ctxt)) of
NONE => raise REFUTE ("print",
"no printer for type " ^ quote (Syntax.string_of_typ ctxt T))
| SOME x => x;
fun print_model ctxt model assignment =
let
val (typs, terms) = model
val typs_msg =
if null typs then
"empty universe (no type variables in term)\n"
else
"Size of types: " ^ commas (map (fn (T, i) =>
Syntax.string_of_typ ctxt T ^ ": " ^ string_of_int i) typs) ^ "\n"
val show_consts_msg =
if not (Config.get ctxt show_consts) andalso Library.exists (is_Const o fst) terms then
"enable \"show_consts\" to show the interpretation of constants\n"
else
""
val terms_msg =
if null terms then
"empty interpretation (no free variables in term)\n"
else
cat_lines (map_filter (fn (t, intr) =>
if Config.get ctxt show_consts orelse not (is_Const t) then
SOME (Syntax.string_of_term ctxt t ^ ": " ^
Syntax.string_of_term ctxt
(print ctxt model (Term.type_of t) intr assignment))
else
NONE) terms) ^ "\n"
in
typs_msg ^ show_consts_msg ^ terms_msg
end;
fun add_interpreter name f = Data.map (fn {interpreters, printers, parameters} =>
case AList.lookup (op =) interpreters name of
NONE => {interpreters = (name, f) :: interpreters,
printers = printers, parameters = parameters}
| SOME _ => error ("Interpreter " ^ name ^ " already declared"));
fun add_printer name f = Data.map (fn {interpreters, printers, parameters} =>
case AList.lookup (op =) printers name of
NONE => {interpreters = interpreters,
printers = (name, f) :: printers, parameters = parameters}
| SOME _ => error ("Printer " ^ name ^ " already declared"));
fun set_default_param (name, value) = Data.map
(fn {interpreters, printers, parameters} =>
{interpreters = interpreters, printers = printers,
parameters = Symtab.update (name, value) parameters});
val get_default_param = Symtab.lookup o #parameters o get_data;
val get_default_params = Symtab.dest o #parameters o get_data;
fun actual_params ctxt override =
let
fun read_bool (parms, name) =
case AList.lookup (op =) parms name of
SOME "true" => true
| SOME "false" => false
| SOME s => error ("parameter " ^ quote name ^
" (value is " ^ quote s ^ ") must be \"true\" or \"false\"")
| NONE => error ("parameter " ^ quote name ^
" must be assigned a value")
fun read_int (parms, name) =
case AList.lookup (op =) parms name of
SOME s =>
(case Int.fromString s of
SOME i => i
| NONE => error ("parameter " ^ quote name ^
" (value is " ^ quote s ^ ") must be an integer value"))
| NONE => error ("parameter " ^ quote name ^
" must be assigned a value")
fun read_string (parms, name) =
case AList.lookup (op =) parms name of
SOME s => s
| NONE => error ("parameter " ^ quote name ^
" must be assigned a value")
val allparams = override @ get_default_params ctxt
val minsize = read_int (allparams, "minsize")
val maxsize = read_int (allparams, "maxsize")
val maxvars = read_int (allparams, "maxvars")
val maxtime = read_int (allparams, "maxtime")
val satsolver = read_string (allparams, "satsolver")
val no_assms = read_bool (allparams, "no_assms")
val expect = the_default "" (AList.lookup (op =) allparams "expect")
val sizes = map_filter
(fn (name, value) => Option.map (pair name) (Int.fromString value))
(filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"
andalso name<>"maxvars" andalso name<>"maxtime"
andalso name<>"satsolver" andalso name<>"no_assms") allparams)
in
{sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars,
maxtime=maxtime, satsolver=satsolver, no_assms=no_assms, expect=expect}
end;
fun typ_of_dtyp _ typ_assoc (Old_Datatype_Aux.DtTFree a) =
the (AList.lookup (op =) typ_assoc (Old_Datatype_Aux.DtTFree a))
| typ_of_dtyp descr typ_assoc (Old_Datatype_Aux.DtType (s, Us)) =
Type (s, map (typ_of_dtyp descr typ_assoc) Us)
| typ_of_dtyp descr typ_assoc (Old_Datatype_Aux.DtRec i) =
let val (s, ds, _) = the (AList.lookup (op =) descr i) in
Type (s, map (typ_of_dtyp descr typ_assoc) ds)
end
val close_form = ATP_Util.close_form
val specialize_type = ATP_Util.specialize_type
fun is_const_of_class thy (s, _) =
let
val class_const_names = map Logic.const_of_class (Sign.all_classes thy)
in
member (op =) class_const_names s
end;
fun is_IDT_constructor thy (s, T) =
(case body_type T of
Type (s', _) =>
(case BNF_LFP_Compat.get_constrs thy s' of
SOME constrs =>
List.exists (fn (cname, cty) =>
cname = s andalso Sign.typ_instance thy (T, cty)) constrs
| NONE => false)
| _ => false);
fun is_IDT_recursor thy (s, _) =
let
val rec_names = Symtab.fold (append o #rec_names o snd) (BNF_LFP_Compat.get_all thy []) []
in
member (op =) rec_names s
end;
fun norm_rhs eqn =
let
fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))
| lambda v t = raise TERM ("lambda", [v, t])
val (lhs, rhs) = Logic.dest_equals eqn
val (_, args) = Term.strip_comb lhs
in
fold lambda (rev args) rhs
end
fun get_def thy (s, T) =
let
fun get_def_ax [] = NONE
| get_def_ax ((axname, ax) :: axioms) =
(let
val (lhs, _) = Logic.dest_equals ax
val c = Term.head_of lhs
val (s', T') = Term.dest_Const c
in
if s=s' then
let
val typeSubs = Sign.typ_match thy (T', T) Vartab.empty
val ax' = Envir.subst_term_types typeSubs ax
val rhs = norm_rhs ax'
in
SOME (axname, rhs)
end
else
get_def_ax axioms
end handle ERROR _ => get_def_ax axioms
| TERM _ => get_def_ax axioms
| Type.TYPE_MATCH => get_def_ax axioms)
in
get_def_ax (Theory.all_axioms_of thy)
end;
fun get_typedef thy T =
let
fun get_typedef_ax [] = NONE
| get_typedef_ax ((axname, ax) :: axioms) =
(let
fun type_of_type_definition (Const (s', T')) =
if s'= \<^const_name>‹type_definition› then
SOME T'
else
NONE
| type_of_type_definition (Free _) = NONE
| type_of_type_definition (Var _) = NONE
| type_of_type_definition (Bound _) = NONE
| type_of_type_definition (Abs (_, _, body)) =
type_of_type_definition body
| type_of_type_definition (t1 $ t2) =
(case type_of_type_definition t1 of
SOME x => SOME x
| NONE => type_of_type_definition t2)
in
case type_of_type_definition ax of
SOME T' =>
let
val T'' = domain_type (domain_type T')
val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty
in
SOME (axname, Envir.subst_term_types typeSubs ax)
end
| NONE => get_typedef_ax axioms
end handle ERROR _ => get_typedef_ax axioms
| TERM _ => get_typedef_ax axioms
| Type.TYPE_MATCH => get_typedef_ax axioms)
in
get_typedef_ax (Theory.all_axioms_of thy)
end;
fun get_classdef thy class =
let
val axname = class ^ "_class_def"
in
Option.map (pair axname)
(AList.lookup (op =) (Theory.all_axioms_of thy) axname)
end;
fun unfold_defs thy t =
let
fun unfold_loop t =
case t of
Const (\<^const_name>‹Pure.all›, _) => t
| Const (\<^const_name>‹Pure.eq›, _) => t
| Const (\<^const_name>‹Pure.imp›, _) => t
| Const (\<^const_name>‹Pure.type›, _) => t
| Const (\<^const_name>‹Trueprop›, _) => t
| Const (\<^const_name>‹Not›, _) => t
|
Const (\<^const_name>‹True›, _) => t
|
Const (\<^const_name>‹False›, _) => t
| Const (\<^const_name>‹undefined›, _) => t
| Const (\<^const_name>‹The›, _) => t
| Const (\<^const_name>‹Hilbert_Choice.Eps›, _) => t
| Const (\<^const_name>‹All›, _) => t
| Const (\<^const_name>‹Ex›, _) => t
| Const (\<^const_name>‹HOL.eq›, _) => t
| Const (\<^const_name>‹HOL.conj›, _) => t
| Const (\<^const_name>‹HOL.disj›, _) => t
| Const (\<^const_name>‹HOL.implies›, _) => t
| Const (\<^const_name>‹Collect›, _) => t
| Const (\<^const_name>‹Set.member›, _) => t
| Const (\<^const_name>‹Finite_Set.card›, _) => t
| Const (\<^const_name>‹Finite_Set.finite›, _) => t
| Const (\<^const_name>‹Orderings.less›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹bool›])])) => t
| Const (\<^const_name>‹Groups.plus›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) => t
| Const (\<^const_name>‹Groups.minus›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) => t
| Const (\<^const_name>‹Groups.times›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) => t
| Const (\<^const_name>‹append›, _) => t
| Const (\<^const_name>‹fst›, _) => t
| Const (\<^const_name>‹snd›, _) => t
| Const (s, T) =>
(if is_IDT_constructor thy (s, T)
orelse is_IDT_recursor thy (s, T) then
t
else
case get_def thy (s, T) of
SOME ( _, rhs) =>
(
unfold_loop rhs)
| NONE => t)
| Free _ => t
| Var _ => t
| Bound _ => t
| Abs (s, T, body) => Abs (s, T, unfold_loop body)
| t1 $ t2 => (unfold_loop t1) $ (unfold_loop t2)
val result = Envir.beta_eta_contract (unfold_loop t)
in
result
end;
local
fun get_axiom thy xname =
Name_Space.check (Context.Theory thy) (Theory.axiom_table thy) (xname, Position.none);
val the_eq_trivial = get_axiom \<^theory>‹HOL› "the_eq_trivial";
val someI = get_axiom \<^theory>‹Hilbert_Choice› "someI";
in
fun collect_axioms ctxt t =
let
val thy = Proof_Context.theory_of ctxt
val _ = tracing "Adding axioms..."
fun collect_this_axiom (axname, ax) axs =
let
val ax' = unfold_defs thy ax
in
if member (op aconv) axs ax' then axs
else (tracing axname; collect_term_axioms ax' (ax' :: axs))
end
and collect_sort_axioms T axs =
let
val sort =
(case T of
TFree (_, sort) => sort
| TVar (_, sort) => sort
| _ => raise REFUTE ("collect_axioms",
"type " ^ Syntax.string_of_typ ctxt T ^ " is not a variable"))
val superclasses = sort @ maps (Sign.super_classes thy) sort
val superclasses = distinct (op =) superclasses
val class_axioms = maps (fn class => map (fn ax =>
("<" ^ class ^ ">", Thm.prop_of ax))
(#axioms (Axclass.get_info thy class) handle ERROR _ => []))
superclasses
val monomorphic_class_axioms = map (fn (axname, ax) =>
(case Term.add_tvars ax [] of
[] => (axname, ax)
| [(idx, S)] => (axname, Envir.subst_term_types (Vartab.make [(idx, (S, T))]) ax)
| _ =>
raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
Syntax.string_of_term ctxt ax ^
") contains more than one type variable")))
class_axioms
in
fold collect_this_axiom monomorphic_class_axioms axs
end
and collect_type_axioms T axs =
case T of
Type (\<^type_name>‹prop›, []) => axs
| Type (\<^type_name>‹fun›, [T1, T2]) => collect_type_axioms T2 (collect_type_axioms T1 axs)
| Type (\<^type_name>‹set›, [T1]) => collect_type_axioms T1 axs
| Type (\<^type_name>‹itself›, [T1]) => collect_type_axioms T1 axs
| Type (s, Ts) =>
(case BNF_LFP_Compat.get_info thy [] s of
SOME _ =>
fold collect_type_axioms Ts axs
| NONE =>
(case get_typedef thy T of
SOME (axname, ax) =>
collect_this_axiom (axname, ax) axs
| NONE =>
fold collect_type_axioms Ts axs))
| TFree _ => collect_sort_axioms T axs
| TVar _ => collect_sort_axioms T axs
and collect_term_axioms t axs =
case t of
Const (\<^const_name>‹Pure.all›, _) => axs
| Const (\<^const_name>‹Pure.eq›, _) => axs
| Const (\<^const_name>‹Pure.imp›, _) => axs
| Const (\<^const_name>‹Pure.type›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹Trueprop›, _) => axs
| Const (\<^const_name>‹Not›, _) => axs
| Const (\<^const_name>‹True›, _) => axs
| Const (\<^const_name>‹False›, _) => axs
| Const (\<^const_name>‹undefined›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹The›, T) =>
let
val ax = specialize_type thy (\<^const_name>‹The›, T) (#2 the_eq_trivial)
in
collect_this_axiom (#1 the_eq_trivial, ax) axs
end
| Const (\<^const_name>‹Hilbert_Choice.Eps›, T) =>
let
val ax = specialize_type thy (\<^const_name>‹Hilbert_Choice.Eps›, T) (#2 someI)
in
collect_this_axiom (#1 someI, ax) axs
end
| Const (\<^const_name>‹All›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹Ex›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹HOL.eq›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹HOL.conj›, _) => axs
| Const (\<^const_name>‹HOL.disj›, _) => axs
| Const (\<^const_name>‹HOL.implies›, _) => axs
| Const (\<^const_name>‹Collect›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹Set.member›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹Finite_Set.card›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹Finite_Set.finite›, T) =>
collect_type_axioms T axs
| Const (\<^const_name>‹Orderings.less›, T as Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹bool›])])) =>
collect_type_axioms T axs
| Const (\<^const_name>‹Groups.plus›, T as Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) =>
collect_type_axioms T axs
| Const (\<^const_name>‹Groups.minus›, T as Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) =>
collect_type_axioms T axs
| Const (\<^const_name>‹Groups.times›, T as Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) =>
collect_type_axioms T axs
| Const (\<^const_name>‹append›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹fst›, T) => collect_type_axioms T axs
| Const (\<^const_name>‹snd›, T) => collect_type_axioms T axs
| Const (s, T) =>
if is_const_of_class thy (s, T) then
let
val class = Logic.class_of_const s
val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), class)
val ax_in = SOME (specialize_type thy (s, T) of_class)
handle Type.TYPE_MATCH => NONE
val ax_1 = Option.map (fn ax => (Syntax.string_of_term ctxt ax, ax)) ax_in
val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) (get_classdef thy class)
in
collect_type_axioms T (fold collect_this_axiom (map_filter I [ax_1, ax_2]) axs)
end
else if is_IDT_constructor thy (s, T)
orelse is_IDT_recursor thy (s, T)
then
collect_type_axioms T axs
else
collect_type_axioms T axs
| Free (_, T) => collect_type_axioms T axs
| Var (_, T) => collect_type_axioms T axs
| Bound _ => axs
| Abs (_, T, body) => collect_term_axioms body (collect_type_axioms T axs)
| t1 $ t2 => collect_term_axioms t2 (collect_term_axioms t1 axs)
val result = map close_form (collect_term_axioms t [])
val _ = tracing " ...done."
in
result
end;
end;
fun ground_types ctxt t =
let
val thy = Proof_Context.theory_of ctxt
fun collect_types T acc =
(case T of
Type (\<^type_name>‹fun›, [T1, T2]) => collect_types T1 (collect_types T2 acc)
| Type (\<^type_name>‹prop›, []) => acc
| Type (\<^type_name>‹set›, [T1]) => collect_types T1 acc
| Type (s, Ts) =>
(case BNF_LFP_Compat.get_info thy [] s of
SOME info =>
let
val index = #index info
val descr = #descr info
val (_, typs, _) = the (AList.lookup (op =) descr index)
val typ_assoc = typs ~~ Ts
val _ = if Library.exists (fn d =>
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) typs then
raise REFUTE ("ground_types", "datatype argument (for type "
^ Syntax.string_of_typ ctxt T ^ ") is not a variable")
else ()
fun collect_dtyp d acc =
let
val dT = typ_of_dtyp descr typ_assoc d
in
case d of
Old_Datatype_Aux.DtTFree _ =>
collect_types dT acc
| Old_Datatype_Aux.DtType (_, ds) =>
collect_types dT (fold_rev collect_dtyp ds acc)
| Old_Datatype_Aux.DtRec i =>
if member (op =) acc dT then
acc
else
let
val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i)
val acc_dT = if Library.exists (fn (_, ds) =>
Library.exists Old_Datatype_Aux.is_rec_type ds) dconstrs then
insert (op =) dT acc
else acc
val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT
val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps
in
acc_dconstrs
end
end
in
collect_dtyp (Old_Datatype_Aux.DtRec index) acc
end
| NONE =>
insert (op =) T (fold collect_types Ts acc))
| TFree _ => insert (op =) T acc
| TVar _ => insert (op =) T acc)
in
fold_types collect_types t []
end;
fun string_of_typ (Type (s, _)) = s
| string_of_typ (TFree (s, _)) = s
| string_of_typ (TVar ((s,_), _)) = s;
fun first_universe xs sizes minsize =
let
fun size_of_typ T =
case AList.lookup (op =) sizes (string_of_typ T) of
SOME n => n
| NONE => minsize
in
map (fn T => (T, size_of_typ T)) xs
end;
fun next_universe xs sizes minsize maxsize =
let
fun make_first _ 0 sum =
if sum = 0 then
SOME []
else
NONE
| make_first max len sum =
if sum <= max orelse max < 0 then
Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0)
else
Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max))
fun next _ _ _ [] =
NONE
| next max len sum [x] =
make_first max (len+1) (sum+x+1)
| next max len sum (x1::x2::xs) =
if x1>0 andalso (x2<max orelse max<0) then
SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
else
next max (len+1) (sum+x1) (x2::xs)
val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs
val diffs = map (fn (_, n) => n-minsize) mutables
in
case next (maxsize-minsize) 0 0 diffs of
SOME diffs' =>
SOME (fst (fold_map (fn (T, _) => fn ds =>
case AList.lookup (op =) sizes (string_of_typ T) of
SOME n => ((T, n), ds)
| NONE => ((T, minsize + hd ds), tl ds))
xs diffs'))
| NONE => NONE
end;
fun toTrue (Leaf [fm, _]) = fm
| toTrue _ = raise REFUTE ("toTrue", "interpretation does not denote a Boolean value");
fun toFalse (Leaf [_, fm]) = fm
| toFalse _ = raise REFUTE ("toFalse", "interpretation does not denote a Boolean value");
fun find_model ctxt
{sizes, minsize, maxsize, maxvars, maxtime, satsolver, no_assms, expect}
assm_ts t negate =
let
val thy = Proof_Context.theory_of ctxt
fun check_expect outcome_code =
if expect = "" orelse outcome_code = expect then outcome_code
else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
fun wrapper () =
let
val time_start = Time.now ()
val t =
if no_assms then t
else if negate then Logic.list_implies (assm_ts, t)
else Logic.mk_conjunction_list (t :: assm_ts)
val u = unfold_defs thy t
val _ = tracing ("Unfolded term: " ^ Syntax.string_of_term ctxt u)
val axioms = collect_axioms ctxt u
val types = fold (union (op =) o ground_types ctxt) (u :: axioms) []
val _ = tracing ("Ground types: "
^ (if null types then "none."
else commas (map (Syntax.string_of_typ ctxt) types)))
val maybe_spurious = Library.exists (fn
Type (s, _) =>
(case BNF_LFP_Compat.get_info thy [] s of
SOME info =>
let
val index = #index info
val descr = #descr info
val (_, _, constrs) = the (AList.lookup (op =) descr index)
in
Library.exists (fn (_, ds) =>
Library.exists Old_Datatype_Aux.is_rec_type ds) constrs
end
| NONE => false)
| _ => false) types
val _ =
if maybe_spurious andalso Context_Position.is_visible ctxt then
warning "Term contains a recursive datatype; countermodel(s) may be spurious!"
else ()
fun find_model_loop universe =
let
val time_spent = Time.now () - time_start
val _ = maxtime = 0 orelse time_spent < Timeout.scale_time (Time.fromSeconds maxtime)
orelse raise Timeout.TIMEOUT time_spent
val init_model = (universe, [])
val init_args = {maxvars = maxvars, def_eq = false, next_idx = 1,
bounds = [], wellformed = True}
val _ = tracing ("Translating term (sizes: "
^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")
val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) =>
let
val (i, m', a') = interpret ctxt m a t'
in
(i, (m', {maxvars = #maxvars a', def_eq = true,
next_idx = #next_idx a', bounds = #bounds a',
wellformed = #wellformed a'}))
end) (u :: axioms) (init_model, init_args)
val fm_u = (if negate then toFalse else toTrue) (hd intrs)
val fm_ax = Prop_Logic.all (map toTrue (tl intrs))
val fm = Prop_Logic.all [#wellformed args, fm_ax, fm_u]
val _ =
(if member (op =) ["cdclite"] satsolver andalso Context_Position.is_visible ctxt then
warning ("Using SAT solver " ^ quote satsolver ^
"; for better performance, consider installing an \
\external solver.")
else ());
val solver =
SAT_Solver.invoke_solver satsolver
handle Option.Option =>
error ("Unknown SAT solver: " ^ quote satsolver ^
". Available solvers: " ^
commas (map (quote o fst) (SAT_Solver.get_solvers ())) ^ ".")
in
writeln "Invoking SAT solver...";
(case solver fm of
SAT_Solver.SATISFIABLE assignment =>
(writeln ("Model found:\n" ^ print_model ctxt model
(fn i => case assignment i of SOME b => b | NONE => true));
if maybe_spurious then "potential" else "genuine")
| SAT_Solver.UNSATISFIABLE _ =>
(writeln "No model exists.";
case next_universe universe sizes minsize maxsize of
SOME universe' => find_model_loop universe'
| NONE => (writeln
"Search terminated, no larger universe within the given limits.";
"none"))
| SAT_Solver.UNKNOWN =>
(writeln "No model found.";
case next_universe universe sizes minsize maxsize of
SOME universe' => find_model_loop universe'
| NONE => (writeln
"Search terminated, no larger universe within the given limits.";
"unknown"))) handle SAT_Solver.NOT_CONFIGURED =>
(error ("SAT solver " ^ quote satsolver ^ " is not configured.");
"unknown")
end
handle MAXVARS_EXCEEDED =>
(writeln ("Search terminated, number of Boolean variables ("
^ string_of_int maxvars ^ " allowed) exceeded.");
"unknown")
val outcome_code = find_model_loop (first_universe types sizes minsize)
in
check_expect outcome_code
end
in
minsize>=1 orelse
error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
maxsize>=1 orelse
error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
maxsize>=minsize orelse
error ("\"maxsize\" (=" ^ string_of_int maxsize ^
") is less than \"minsize\" (=" ^ string_of_int minsize ^ ").");
maxvars>=0 orelse
error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
maxtime>=0 orelse
error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
writeln ("Trying to find a model that "
^ (if negate then "refutes" else "satisfies") ^ ": "
^ Syntax.string_of_term ctxt t);
Timeout.apply (Time.fromSeconds maxtime)
wrapper ()
handle Timeout.TIMEOUT _ =>
(writeln ("Search terminated, time limit (" ^
string_of_int maxtime
^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.");
check_expect "unknown")
end;
fun satisfy_term ctxt params assm_ts t =
find_model ctxt (actual_params ctxt params) assm_ts t false;
fun refute_term ctxt params assm_ts t =
let
val _ = null (Term.add_tvars t []) orelse
error "Term to be refuted contains schematic type variables"
val vars = sort_by (fst o fst) (Term.add_vars t [])
val ex_closure = fold (fn ((x, i), T) => fn t' =>
HOLogic.exists_const T $
Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
fun strip_all_body (Const (\<^const_name>‹Pure.all›, _) $ Abs (_, _, t)) =
strip_all_body t
| strip_all_body (Const (\<^const_name>‹Trueprop›, _) $ t) =
strip_all_body t
| strip_all_body (Const (\<^const_name>‹All›, _) $ Abs (_, _, t)) =
strip_all_body t
| strip_all_body t = t
fun strip_all_vars (Const (\<^const_name>‹Pure.all›, _) $ Abs (a, T, t)) =
(a, T) :: strip_all_vars t
| strip_all_vars (Const (\<^const_name>‹Trueprop›, _) $ t) =
strip_all_vars t
| strip_all_vars (Const (\<^const_name>‹All›, _) $ Abs (a, T, t)) =
(a, T) :: strip_all_vars t
| strip_all_vars _ = [] : (string * typ) list
val strip_t = strip_all_body ex_closure
val frees = Term.rename_wrt_term strip_t (strip_all_vars ex_closure)
val subst_t = Term.subst_bounds (map Free frees, strip_t)
in
find_model ctxt (actual_params ctxt params) assm_ts subst_t true
end;
fun refute_goal ctxt params th i =
let
val t = th |> Thm.prop_of
in
if Logic.no_prems t then
(writeln "No subgoal!"; "none")
else
let
val assms = map Thm.term_of (Assumption.all_assms_of ctxt)
val (t, frees) = Logic.goal_params t i
in
refute_term ctxt params assms (subst_bounds (frees, t))
end
end
fun make_constants ctxt model T =
let
fun unit_vectors n =
let
fun unit_vector (k, n) =
Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
fun unit_vectors_loop k =
if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1)
in
unit_vectors_loop 1
end
fun pick_all 1 xs = map single xs
| pick_all n xs =
let val rec_pick = pick_all (n - 1) xs in
maps (fn x => map (cons x) rec_pick) xs
end
fun make_constants_intr (Leaf xs) = unit_vectors (length xs)
| make_constants_intr (Node xs) = map Node (pick_all (length xs)
(make_constants_intr (hd xs)))
val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1,
bounds=[], wellformed=True} (Free ("dummy", T))
in
make_constants_intr i
end;
fun size_of_type ctxt model T =
let
fun size_of_intr (Leaf xs) = length xs
| size_of_intr (Node xs) = Integer.pow (length xs) (size_of_intr (hd xs))
val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1,
bounds=[], wellformed=True} (Free ("dummy", T))
in
size_of_intr i
end;
val TT = Leaf [True, False];
val FF = Leaf [False, True];
fun make_equality (i1, i2) =
let
fun equal (i1, i2) =
(case i1 of
Leaf xs =>
(case i2 of
Leaf ys => Prop_Logic.dot_product (xs, ys)
| Node _ => raise REFUTE ("make_equality",
"second interpretation is higher"))
| Node xs =>
(case i2 of
Leaf _ => raise REFUTE ("make_equality",
"first interpretation is higher")
| Node ys => Prop_Logic.all (map equal (xs ~~ ys))))
fun not_equal (i1, i2) =
(case i1 of
Leaf xs =>
(case i2 of
Leaf ys => Prop_Logic.all ((Prop_Logic.exists xs)
:: (Prop_Logic.exists ys)
:: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys)))
| Node _ => raise REFUTE ("make_equality",
"second interpretation is higher"))
| Node xs =>
(case i2 of
Leaf _ => raise REFUTE ("make_equality",
"first interpretation is higher")
| Node ys => Prop_Logic.exists (map not_equal (xs ~~ ys))))
in
Leaf [equal (i1, i2), not_equal (i1, i2)]
end;
fun make_def_equality (i1, i2) =
let
fun equal (i1, i2) =
(case i1 of
Leaf xs =>
(case i2 of
Leaf ys => SOr (Prop_Logic.dot_product (xs, ys),
SAnd (Prop_Logic.all (map SNot xs), Prop_Logic.all (map SNot ys)))
| Node _ => raise REFUTE ("make_def_equality",
"second interpretation is higher"))
| Node xs =>
(case i2 of
Leaf _ => raise REFUTE ("make_def_equality",
"first interpretation is higher")
| Node ys => Prop_Logic.all (map equal (xs ~~ ys))))
val eq = equal (i1, i2)
in
Leaf [eq, SNot eq]
end;
fun interpretation_apply (i1, i2) =
let
fun interpretation_disjunction (tr1,tr2) =
tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys))
(tree_pair (tr1,tr2))
fun prop_formula_times_interpretation (fm,tr) =
tree_map (map (fn x => SAnd (fm,x))) tr
fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =
prop_formula_times_interpretation (fm,tr)
| prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =
interpretation_disjunction (prop_formula_times_interpretation (fm,tr),
prop_formula_list_dot_product_interpretation_list (fms,trees))
| prop_formula_list_dot_product_interpretation_list (_,_) =
raise REFUTE ("interpretation_apply", "empty list (in dot product)")
fun pick_all [xs] = map single xs
| pick_all (xs::xss) =
let val rec_pick = pick_all xss in
maps (fn x => map (cons x) rec_pick) xs
end
| pick_all _ = raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
fun interpretation_to_prop_formula_list (Leaf xs) = xs
| interpretation_to_prop_formula_list (Node trees) =
map Prop_Logic.all (pick_all
(map interpretation_to_prop_formula_list trees))
in
case i1 of
Leaf _ =>
raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
| Node xs =>
prop_formula_list_dot_product_interpretation_list
(interpretation_to_prop_formula_list i2, xs)
end;
fun eta_expand t i =
let
val Ts = Term.binder_types (Term.fastype_of t)
val t' = Term.incr_boundvars i t
in
fold_rev (fn T => fn term => Abs ("<eta_expand>", T, term))
(List.take (Ts, i))
(Term.list_comb (t', map Bound (i-1 downto 0)))
end;
fun size_of_dtyp ctxt typ_sizes descr typ_assoc constructors =
Integer.sum (map (fn (_, dtyps) =>
Integer.prod (map (size_of_type ctxt (typ_sizes, []) o
(typ_of_dtyp descr typ_assoc)) dtyps))
constructors);
fun stlc_interpreter ctxt model args t =
let
val (typs, terms) = model
val {maxvars, def_eq, next_idx, bounds, wellformed} = args
fun interpret_groundterm T =
let
fun interpret_groundtype () =
let
val size =
if T = Term.propT then 2
else the (AList.lookup (op =) typs T)
val next = next_idx + size
val _ = (if next - 1 > maxvars andalso maxvars > 0 then
raise MAXVARS_EXCEEDED else ())
val fms = map BoolVar (next_idx upto (next_idx + size - 1))
val intr = Leaf fms
fun one_of_two_false [] = True
| one_of_two_false (x::xs) = SAnd (Prop_Logic.all (map (fn x' =>
SOr (SNot x, SNot x')) xs), one_of_two_false xs)
val wf = one_of_two_false fms
in
SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
def_eq = def_eq, next_idx = next, bounds = bounds,
wellformed = SAnd (wellformed, wf)})
end
in
case T of
Type ("fun", [T1, T2]) =>
let
fun make_copies idx 0 = (idx, [], True)
| make_copies idx n =
let
val (copy, _, new_args) = interpret ctxt (typs, [])
{maxvars = maxvars, def_eq = false, next_idx = idx,
bounds = [], wellformed = True} (Free ("dummy", T2))
val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1)
in
(idx', copy :: copies, SAnd (#wellformed new_args, wf'))
end
val (next, copies, wf) = make_copies next_idx
(size_of_type ctxt model T1)
val intr = Node copies
in
SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
def_eq = def_eq, next_idx = next, bounds = bounds,
wellformed = SAnd (wellformed, wf)})
end
| Type _ => interpret_groundtype ()
| TFree _ => interpret_groundtype ()
| TVar _ => interpret_groundtype ()
end
in
case AList.lookup (op =) terms t of
SOME intr =>
SOME (intr, model, args)
| NONE =>
(case t of
Const (_, T) => interpret_groundterm T
| Free (_, T) => interpret_groundterm T
| Var (_, T) => interpret_groundterm T
| Bound i => SOME (nth (#bounds args) i, model, args)
| Abs (_, T, body) =>
let
val constants = make_constants ctxt model T
val (bodies, (model', args')) = fold_map
(fn c => fn (m, a) =>
let
val (i', m', a') = interpret ctxt m {maxvars = #maxvars a,
def_eq = #def_eq a, next_idx = #next_idx a,
bounds = (c :: #bounds a), wellformed = #wellformed a} body
in
(i', (m', {maxvars = maxvars, def_eq = def_eq,
next_idx = #next_idx a', bounds = bounds,
wellformed = #wellformed a'}))
end)
constants (model, args)
in
SOME (Node bodies, model', args')
end
| t1 $ t2 =>
let
val (intr1, model1, args1) = interpret ctxt model args t1
val (intr2, model2, args2) = interpret ctxt model1 args1 t2
in
SOME (interpretation_apply (intr1, intr2), model2, args2)
end)
end;
fun Pure_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Pure.all›, _) $ t1 =>
let
val (i, m, a) = interpret ctxt model args t1
in
case i of
Node xs =>
let
val fmTrue = Prop_Logic.all (map toTrue xs)
val fmFalse = Prop_Logic.exists (map toFalse xs)
in
SOME (Leaf [fmTrue, fmFalse], m, a)
end
| _ =>
raise REFUTE ("Pure_interpreter",
"\"Pure.all\" is followed by a non-function")
end
| Const (\<^const_name>‹Pure.all›, _) =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹Pure.eq›, _) $ t1 $ t2 =>
let
val (i1, m1, a1) = interpret ctxt model args t1
val (i2, m2, a2) = interpret ctxt m1 a1 t2
in
SOME ((if #def_eq args then make_def_equality else make_equality)
(i1, i2), m2, a2)
end
| Const (\<^const_name>‹Pure.eq›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹Pure.eq›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| Const (\<^const_name>‹Pure.imp›, _) $ t1 $ t2 =>
let
val (i1, m1, a1) = interpret ctxt model args t1
val (i2, m2, a2) = interpret ctxt m1 a1 t2
val fmTrue = Prop_Logic.SOr (toFalse i1, toTrue i2)
val fmFalse = Prop_Logic.SAnd (toTrue i1, toFalse i2)
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
| Const (\<^const_name>‹Pure.imp›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹Pure.imp›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| _ => NONE;
fun HOLogic_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Trueprop›, _) =>
SOME (Node [TT, FF], model, args)
| Const (\<^const_name>‹Not›, _) =>
SOME (Node [FF, TT], model, args)
| Const (\<^const_name>‹True›, _) =>
SOME (TT, model, args)
| Const (\<^const_name>‹False›, _) =>
SOME (FF, model, args)
| Const (\<^const_name>‹All›, _) $ t1 =>
let
val (i, m, a) = interpret ctxt model args t1
in
case i of
Node xs =>
let
val fmTrue = Prop_Logic.all (map toTrue xs)
val fmFalse = Prop_Logic.exists (map toFalse xs)
in
SOME (Leaf [fmTrue, fmFalse], m, a)
end
| _ =>
raise REFUTE ("HOLogic_interpreter",
"\"All\" is followed by a non-function")
end
| Const (\<^const_name>‹All›, _) =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹Ex›, _) $ t1 =>
let
val (i, m, a) = interpret ctxt model args t1
in
case i of
Node xs =>
let
val fmTrue = Prop_Logic.exists (map toTrue xs)
val fmFalse = Prop_Logic.all (map toFalse xs)
in
SOME (Leaf [fmTrue, fmFalse], m, a)
end
| _ =>
raise REFUTE ("HOLogic_interpreter",
"\"Ex\" is followed by a non-function")
end
| Const (\<^const_name>‹Ex›, _) =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹HOL.eq›, _) $ t1 $ t2 =>
let
val (i1, m1, a1) = interpret ctxt model args t1
val (i2, m2, a2) = interpret ctxt m1 a1 t2
in
SOME (make_equality (i1, i2), m2, a2)
end
| Const (\<^const_name>‹HOL.eq›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹HOL.eq›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| Const (\<^const_name>‹HOL.conj›, _) $ t1 $ t2 =>
let
val (i1, m1, a1) = interpret ctxt model args t1
val (i2, m2, a2) = interpret ctxt m1 a1 t2
val fmTrue = Prop_Logic.SAnd (toTrue i1, toTrue i2)
val fmFalse = Prop_Logic.SOr (toFalse i1, toFalse i2)
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
| Const (\<^const_name>‹HOL.conj›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹HOL.conj›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| Const (\<^const_name>‹HOL.disj›, _) $ t1 $ t2 =>
let
val (i1, m1, a1) = interpret ctxt model args t1
val (i2, m2, a2) = interpret ctxt m1 a1 t2
val fmTrue = Prop_Logic.SOr (toTrue i1, toTrue i2)
val fmFalse = Prop_Logic.SAnd (toFalse i1, toFalse i2)
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
| Const (\<^const_name>‹HOL.disj›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹HOL.disj›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| Const (\<^const_name>‹HOL.implies›, _) $ t1 $ t2 =>
let
val (i1, m1, a1) = interpret ctxt model args t1
val (i2, m2, a2) = interpret ctxt m1 a1 t2
val fmTrue = Prop_Logic.SOr (toFalse i1, toTrue i2)
val fmFalse = Prop_Logic.SAnd (toTrue i1, toFalse i2)
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
| Const (\<^const_name>‹HOL.implies›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹HOL.implies›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| _ => NONE;
fun IDT_interpreter ctxt model args t =
let
val thy = Proof_Context.theory_of ctxt
val (typs, terms) = model
fun interpret_term (Type (s, Ts)) =
(case BNF_LFP_Compat.get_info thy [] s of
SOME info =>
let
val depth = AList.lookup (op =) typs (Type (s, Ts))
val _ =
(case depth of SOME n =>
if n < 0 then
raise REFUTE ("IDT_interpreter", "negative depth")
else ()
| _ => ())
in
if depth = (SOME 0) then
SOME (Leaf [], model, args)
else
let
val index = #index info
val descr = #descr info
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
val typ_assoc = dtyps ~~ Ts
val _ =
if Library.exists (fn d =>
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_interpreter",
"datatype argument (for type "
^ Syntax.string_of_typ ctxt (Type (s, Ts))
^ ") is not a variable")
else ()
val typs' = case depth of NONE => typs | SOME n =>
AList.update (op =) (Type (s, Ts), n-1) typs
val size = size_of_dtyp ctxt typs' descr typ_assoc constrs
val next_idx = #next_idx args
val next = next_idx+size
val _ = (if next-1 > #maxvars args andalso
#maxvars args > 0 then raise MAXVARS_EXCEEDED else ())
val fms = map BoolVar (next_idx upto (next_idx+size-1))
val intr = Leaf fms
fun one_of_two_false [] = True
| one_of_two_false (x::xs) = SAnd (Prop_Logic.all (map (fn x' =>
SOr (SNot x, SNot x')) xs), one_of_two_false xs)
val wf = one_of_two_false fms
in
SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args,
def_eq = #def_eq args, next_idx = next, bounds = #bounds args,
wellformed = SAnd (#wellformed args, wf)})
end
end
| NONE =>
NONE)
| interpret_term _ =
NONE
in
case AList.lookup (op =) terms t of
SOME intr =>
SOME (intr, model, args)
| NONE =>
(case t of
Free (_, T) => interpret_term T
| Var (_, T) => interpret_term T
| Const (_, T) => interpret_term T
| _ => NONE)
end;
fun IDT_constructor_interpreter ctxt model args t =
let
val thy = Proof_Context.theory_of ctxt
fun canonical_terms typs T =
(case T of
Type ("fun", [T1, T2]) =>
let
fun pick_all 1 xs = map single xs
| pick_all n xs =
let val rec_pick = pick_all (n-1) xs in
maps (fn x => map (cons x) rec_pick) xs
end
val terms1 = canonical_terms typs T1
val terms2 = canonical_terms typs T2
val functions = map (curry (op ~~) terms1)
(pick_all (length terms1) terms2)
val pairss = map (map HOLogic.mk_prod) functions
val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT
val HOLogic_empty_set = Const (\<^const_abbrev>‹Set.empty›, HOLogic_setT)
val HOLogic_insert =
Const (\<^const_name>‹insert›, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
in
map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps
HOLogic_empty_set) pairss
end
| Type (s, Ts) =>
(case BNF_LFP_Compat.get_info thy [] s of
SOME info =>
(case AList.lookup (op =) typs T of
SOME 0 =>
[]
| _ =>
let
val index = #index info
val descr = #descr info
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
val typ_assoc = dtyps ~~ Ts
val _ =
if Library.exists (fn d =>
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_constructor_interpreter",
"datatype argument (for type "
^ Syntax.string_of_typ ctxt T
^ ") is not a variable")
else ()
val typs' =
(case AList.lookup (op =) typs T of NONE => typs
| SOME n => AList.update (op =) (T, n-1) typs)
fun constructor_terms terms [] = terms
| constructor_terms terms (d::ds) =
let
val dT = typ_of_dtyp descr typ_assoc d
val d_terms = canonical_terms typs' dT
in
constructor_terms
(map_product (curry op $) terms d_terms) ds
end
in
maps (fn (cname, ctyps) =>
let
val cTerm = Const (cname,
map (typ_of_dtyp descr typ_assoc) ctyps ---> T)
in
constructor_terms [cTerm] ctyps
end) constrs
end)
| NONE =>
map (fn intr => print ctxt (typs, []) T intr (K false))
(make_constants ctxt (typs, []) T))
| _ =>
map (fn intr => print ctxt (typs, []) T intr (K false))
(make_constants ctxt (typs, []) T))
val (typs, terms) = model
in
case AList.lookup (op =) terms t of
SOME intr =>
SOME (intr, model, args)
| NONE =>
(case t of
Const (s, T) =>
(case body_type T of
Type (s', Ts') =>
(case BNF_LFP_Compat.get_info thy [] s' of
SOME info =>
let
val index = #index info
val descr = #descr info
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
val typ_assoc = dtyps ~~ Ts'
val _ = if Library.exists (fn d =>
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_constructor_interpreter",
"datatype argument (for type "
^ Syntax.string_of_typ ctxt (Type (s', Ts'))
^ ") is not a variable")
else ()
val (constrs1, constrs2) = chop_prefix (fn (cname, ctypes) =>
not (cname = s andalso Sign.typ_instance thy (T,
map (typ_of_dtyp descr typ_assoc) ctypes
---> Type (s', Ts')))) constrs
in
case constrs2 of
[] =>
NONE
| (_, ctypes)::_ =>
let
val depth = AList.lookup (op =) typs (Type (s', Ts'))
val _ = (case depth of SOME 0 =>
raise REFUTE ("IDT_constructor_interpreter",
"depth is 0")
| _ => ())
val typs' = (case depth of NONE => typs | SOME n =>
AList.update (op =) (Type (s', Ts'), n-1) typs)
val offset = size_of_dtyp ctxt typs' descr typ_assoc constrs1
val total = offset +
size_of_dtyp ctxt typs' descr typ_assoc constrs2
val _ = if total <> size_of_type ctxt (typs, [])
(Type (s', Ts')) then
raise REFUTE ("IDT_constructor_interpreter",
"total is not equal to current size")
else ()
fun make_undef [] = Leaf (replicate total False)
| make_undef (d::ds) =
let
val dT = typ_of_dtyp descr typ_assoc d
val size = size_of_type ctxt (typs, []) dT
in
Node (replicate size (make_undef ds))
end
fun make_constr [] offset =
if offset < total then
(Leaf (replicate offset False @ True ::
(replicate (total - offset - 1) False)), offset + 1)
else
raise REFUTE ("IDT_constructor_interpreter",
"offset >= total")
| make_constr (d::ds) offset =
let
val dT = typ_of_dtyp descr typ_assoc d
val terms' = canonical_terms typs' dT
val _ =
if length terms' <> size_of_type ctxt (typs', []) dT
then
raise REFUTE ("IDT_constructor_interpreter",
"length of terms' is not equal to old size")
else ()
val terms = canonical_terms typs dT
val _ =
if length terms <> size_of_type ctxt (typs, []) dT
then
raise REFUTE ("IDT_constructor_interpreter",
"length of terms is not equal to current size")
else ()
val _ =
if length terms < length terms' then
raise REFUTE ("IDT_constructor_interpreter",
"current size is less than old size")
else ()
val _ =
if forall (member (op =) terms) terms' then ()
else
raise REFUTE ("IDT_constructor_interpreter",
"element has disappeared")
val _ =
let
fun search (x::xs) (y::ys) =
if x = y then search xs ys else search (x::xs) ys
| search (_::_) [] =
raise REFUTE ("IDT_constructor_interpreter",
"element order not preserved")
| search [] _ = ()
in search terms' terms end
val (intrs, new_offset) =
fold_map (fn t_elem => fn off =>
if member (op =) terms' t_elem then
make_constr ds off
else
(make_undef ds, off))
terms offset
in
(Node intrs, new_offset)
end
in
SOME (fst (make_constr ctypes offset), model, args)
end
end
| NONE =>
NONE)
| _ =>
NONE)
| _ =>
NONE)
end;
fun IDT_recursion_interpreter ctxt model args t =
let
val thy = Proof_Context.theory_of ctxt
in
case strip_comb t of
(Const (s, T), params) =>
Symtab.fold (fn (_, info) => fn result =>
case result of
SOME _ =>
result
| NONE =>
if member (op =) (#rec_names info) s then
let
val mconstrs_count =
Integer.sum (map (fn (_, (_, _, cs)) => length cs) (#descr info))
in
if mconstrs_count < length params then
NONE
else if mconstrs_count > length params then
SOME (interpret ctxt model args (eta_expand t
(mconstrs_count - length params)))
else
let
val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) =>
let
val (i, m', a') = interpret ctxt m a p
in
(i, (m', a'))
end) params (model, args)
val (typs, _) = model'
val index = #index info
val descr = #descr info
val (_, dtyps, _) = the (AList.lookup (op =) descr index)
val _ = if map fst descr <> (0 upto (length descr - 1)) then
raise REFUTE ("IDT_recursion_interpreter",
"order of constructors and corresponding parameters/" ^
"recursion operators and corresponding datatypes " ^
"different?")
else ()
val _ =
if Library.exists (fn d =>
case d of Old_Datatype_Aux.DtTFree _ => false
| _ => true) dtyps
then
raise REFUTE ("IDT_recursion_interpreter",
"datatype argument is not a variable")
else ()
val IDT = nth (binder_types T) mconstrs_count
val idt_index = find_index (fn s' => s' = s) (#rec_names info)
fun rec_typ_assoc acc [] = acc
| rec_typ_assoc acc ((d, T)::xs) =
(case AList.lookup op= acc d of
NONE =>
(case d of
Old_Datatype_Aux.DtTFree _ =>
rec_typ_assoc ((d, T)::acc) xs
| Old_Datatype_Aux.DtType (s, ds) =>
let
val (s', Ts) = dest_Type T
in
if s=s' then
rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
else
raise REFUTE ("IDT_recursion_interpreter",
"DtType/Type mismatch")
end
| Old_Datatype_Aux.DtRec i =>
let
val (_, ds, _) = the (AList.lookup (op =) descr i)
val (_, Ts) = dest_Type T
in
rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
end)
| SOME T' =>
if T=T' then
rec_typ_assoc acc xs
else
raise REFUTE ("IDT_recursion_interpreter",
"different type associations for the same dtyp"))
val typ_assoc = filter
(fn (Old_Datatype_Aux.DtTFree _, _) => true | (_, _) => false)
(rec_typ_assoc []
(#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))
val _ =
if not (eq_set (op =) (dtyps, map fst typ_assoc))
then
raise REFUTE ("IDT_recursion_interpreter",
"type association has extra/missing elements")
else ()
val mc_intrs = map (fn (idx, (_, _, cs)) =>
let
val c_return_typ = typ_of_dtyp descr typ_assoc
(Old_Datatype_Aux.DtRec idx)
in
(idx, map (fn (cname, cargs) =>
(#1 o interpret ctxt (typs, []) {maxvars=0,
def_eq=false, next_idx=1, bounds=[],
wellformed=True}) (Const (cname, map (typ_of_dtyp
descr typ_assoc) cargs ---> c_return_typ))) cs)
end) descr
val (mc_p_intrs, p_intrs') = fold_map
(fn (idx, c_intrs) => fn p_intrs' =>
let
val len = length c_intrs
in
((idx, c_intrs ~~ List.take (p_intrs', len)),
List.drop (p_intrs', len))
end) mc_intrs p_intrs
val _ =
if p_intrs' <> [] then
raise REFUTE ("IDT_recursion_interpreter",
"more parameter than constructor interpretations")
else ()
fun ci_pi (Leaf xs, pi) =
if List.exists (equal True) xs then [pi] else []
| ci_pi (Node xs, Node ys) = maps ci_pi (xs ~~ ys)
| ci_pi (Node _, Leaf _) =
raise REFUTE ("IDT_recursion_interpreter",
"constructor takes more arguments than the " ^
"associated parameter")
val rec_operators = map (fn (idx, c_p_intrs) =>
(idx, maps ci_pi c_p_intrs)) mc_p_intrs
val _ = map (fn (idx, intrs) =>
let
val T = typ_of_dtyp descr typ_assoc
(Old_Datatype_Aux.DtRec idx)
in
if length intrs <> size_of_type ctxt (typs, []) T then
raise REFUTE ("IDT_recursion_interpreter",
"wrong number of interpretations for rec. operator")
else ()
end) rec_operators
val REC_OPERATORS = map (fn (idx, intrs) =>
(idx, Array.fromList (map (pair false) intrs)))
rec_operators
fun get_args (Leaf xs) elem =
if find_index (fn x => x = True) xs = elem then
SOME []
else
NONE
| get_args (Node xs) elem =
let
fun search ([], _) =
NONE
| search (x::xs, n) =
(case get_args x elem of
SOME result => SOME (n::result)
| NONE => search (xs, n+1))
in
search (xs, 0)
end
fun get_cargs idx elem =
let
fun get_cargs_rec (_, []) =
raise REFUTE ("IDT_recursion_interpreter",
"no matching constructor found for datatype element")
| get_cargs_rec (n, x::xs) =
(case get_args x elem of
SOME args => (n, args)
| NONE => get_cargs_rec (n+1, xs))
in
get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx))
end
fun compute_array_entry idx elem =
let
val arr = the (AList.lookup (op =) REC_OPERATORS idx)
val (flag, intr) = Array.sub (arr, elem)
in
if flag then
intr
else
let
val (c, args) = get_cargs idx elem
val (_, _, constrs) = the (AList.lookup (op =) descr idx)
val (_, dtyps) = nth constrs c
val rec_dtyps_args = filter
(Old_Datatype_Aux.is_rec_type o fst) (dtyps ~~ args)
val rec_dtyps_intrs = map (fn (dtyp, arg) =>
let
val dT = typ_of_dtyp descr typ_assoc dtyp
val consts = make_constants ctxt (typs, []) dT
val arg_i = nth consts arg
in
(dtyp, arg_i)
end) rec_dtyps_args
fun rec_intr (Old_Datatype_Aux.DtRec i) (Leaf xs) =
compute_array_entry i (find_index (fn x => x = True) xs)
| rec_intr (Old_Datatype_Aux.DtRec _) (Node _) =
raise REFUTE ("IDT_recursion_interpreter",
"interpretation for IDT is a node")
| rec_intr (Old_Datatype_Aux.DtType ("fun", [_, dt2])) (Node xs) =
Node (map (rec_intr dt2) xs)
| rec_intr (Old_Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) =
raise REFUTE ("IDT_recursion_interpreter",
"interpretation for function dtyp is a leaf")
| rec_intr _ _ =
raise REFUTE ("IDT_recursion_interpreter",
"non-recursive codomain in recursive dtyp")
val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs
val result = fold (fn arg_i => fn i =>
interpretation_apply (i, arg_i)) arg_intrs intr
val _ = Array.upd arr elem (true, result)
in
result
end
end
val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index))
val _ =
if idt_size <> size_of_type ctxt (typs, []) IDT then
raise REFUTE ("IDT_recursion_interpreter",
"unexpected size of IDT (wrong type associated?)")
else ()
val rec_op = Node (map_range (compute_array_entry idt_index) idt_size)
in
SOME (rec_op, model', args')
end
end
else
NONE
) (BNF_LFP_Compat.get_all thy []) NONE
| _ =>
NONE
end;
fun set_interpreter ctxt model args t =
let
val (typs, terms) = model
in
case AList.lookup (op =) terms t of
SOME intr =>
SOME (intr, model, args)
| NONE =>
(case t of
Free (x, Type (\<^type_name>‹set›, [T])) =>
let
val (intr, _, args') =
interpret ctxt (typs, []) args (Free (x, T --> HOLogic.boolT))
in
SOME (intr, (typs, (t, intr)::terms), args')
end
| Var ((x, i), Type (\<^type_name>‹set›, [T])) =>
let
val (intr, _, args') =
interpret ctxt (typs, []) args (Var ((x,i), T --> HOLogic.boolT))
in
SOME (intr, (typs, (t, intr)::terms), args')
end
| Const (s, Type (\<^type_name>‹set›, [T])) =>
let
val (intr, _, args') =
interpret ctxt (typs, []) args (Const (s, T --> HOLogic.boolT))
in
SOME (intr, (typs, (t, intr)::terms), args')
end
| Const (\<^const_name>‹Collect›, _) $ t1 =>
SOME (interpret ctxt model args t1)
| Const (\<^const_name>‹Collect›, _) =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹Set.member›, _) $ t1 $ t2 =>
SOME (interpret ctxt model args (t2 $ t1))
| Const (\<^const_name>‹Set.member›, _) $ _ =>
SOME (interpret ctxt model args (eta_expand t 1))
| Const (\<^const_name>‹Set.member›, _) =>
SOME (interpret ctxt model args (eta_expand t 2))
| _ => NONE)
end;
fun Finite_Set_card_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Finite_Set.card›,
Type ("fun", [Type (\<^type_name>‹set›, [T]), \<^typ>‹nat›])) =>
let
fun number_of_elements (Node xs) =
fold (fn x => fn n =>
if x = TT then
n + 1
else if x = FF then
n
else
raise REFUTE ("Finite_Set_card_interpreter",
"interpretation for set type does not yield a Boolean"))
xs 0
| number_of_elements (Leaf _) =
raise REFUTE ("Finite_Set_card_interpreter",
"interpretation for set type is a leaf")
val size_of_nat = size_of_type ctxt model (\<^typ>‹nat›)
fun card i =
let
val n = number_of_elements i
in
if n < size_of_nat then
Leaf ((replicate n False) @ True ::
(replicate (size_of_nat-n-1) False))
else
Leaf (replicate size_of_nat False)
end
val set_constants = make_constants ctxt model (HOLogic.mk_setT T)
in
SOME (Node (map card set_constants), model, args)
end
| _ => NONE;
fun Finite_Set_finite_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Finite_Set.finite›,
Type ("fun", [_, \<^typ>‹bool›])) $ _ =>
SOME (TT, model, args)
| Const (\<^const_name>‹Finite_Set.finite›,
Type ("fun", [set_T, \<^typ>‹bool›])) =>
let
val size_of_set = size_of_type ctxt model set_T
in
SOME (Node (replicate size_of_set TT), model, args)
end
| _ => NONE;
fun Nat_less_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Orderings.less›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹bool›])])) =>
let
val size_of_nat = size_of_type ctxt model (\<^typ>‹nat›)
fun less n = Node ((replicate n FF) @ (replicate (size_of_nat - n) TT))
in
SOME (Node (map less (1 upto size_of_nat)), model, args)
end
| _ => NONE;
fun Nat_plus_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Groups.plus›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) =>
let
val size_of_nat = size_of_type ctxt model (\<^typ>‹nat›)
fun plus m n =
let
val element = m + n
in
if element > size_of_nat - 1 then
Leaf (replicate size_of_nat False)
else
Leaf ((replicate element False) @ True ::
(replicate (size_of_nat - element - 1) False))
end
in
SOME (Node (map_range (fn m => Node (map_range (plus m) size_of_nat)) size_of_nat),
model, args)
end
| _ => NONE;
fun Nat_minus_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Groups.minus›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) =>
let
val size_of_nat = size_of_type ctxt model (\<^typ>‹nat›)
fun minus m n =
let
val element = Int.max (m-n, 0)
in
Leaf ((replicate element False) @ True ::
(replicate (size_of_nat - element - 1) False))
end
in
SOME (Node (map_range (fn m => Node (map_range (minus m) size_of_nat)) size_of_nat),
model, args)
end
| _ => NONE;
fun Nat_times_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹Groups.times›, Type ("fun", [\<^typ>‹nat›,
Type ("fun", [\<^typ>‹nat›, \<^typ>‹nat›])])) =>
let
val size_of_nat = size_of_type ctxt model (\<^typ>‹nat›)
fun mult m n =
let
val element = m * n
in
if element > size_of_nat - 1 then
Leaf (replicate size_of_nat False)
else
Leaf ((replicate element False) @ True ::
(replicate (size_of_nat - element - 1) False))
end
in
SOME (Node (map_range (fn m => Node (map_range (mult m) size_of_nat)) size_of_nat),
model, args)
end
| _ => NONE;
fun List_append_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹append›,
Type (\<^type_name>‹fun›, [Type (\<^type_name>‹list›, [T]),
Type (\<^type_name>‹fun›, [Type (\<^type_name>‹list›, [_]), Type (\<^type_name>‹list›, [_])])])) =>
let
val size_elem = size_of_type ctxt model T
val size_list = size_of_type ctxt model (Type (\<^type_name>‹list›, [T]))
val list_length =
let
fun list_length_acc len lists total =
if lists = total then
len
else if lists < total then
list_length_acc (len+1) (lists*size_elem) (total-lists)
else
raise REFUTE ("List_append_interpreter",
"size_list not equal to 1 + size_elem + ... + " ^
"size_elem^len, for some len")
in
list_length_acc 0 1 size_list
end
val elements = 0 upto (size_list-1)
val (lenoff_lists, _) = fold_map (fn elem => fn (offsets, off) =>
let
val len = length offsets
val assoc = (elem, (len, off))
in
if len < list_length then
(assoc, (off :: offsets, off * size_elem))
else if off mod size_elem < size_elem - 1 then
(assoc, (offsets, off + 1))
else
let
val offsets' = drop_prefix (fn off' => off' mod size_elem = size_elem - 1) offsets
in
case offsets' of
[] =>
(assoc, ([], 0))
| off'::offs' =>
(assoc, (offs', off' + 1))
end
end) elements ([], 0)
val lenoff'_lists = map Library.swap lenoff_lists
fun append m n =
let
val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m)
val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n)
val len_elem = len_m + len_n
val off_elem = off_m * Integer.pow len_n size_elem + off_n
in
case AList.lookup op= lenoff'_lists (len_elem, off_elem) of
NONE =>
Leaf (replicate size_list False)
| SOME element =>
Leaf ((replicate element False) @ True ::
(replicate (size_list - element - 1) False))
end
in
SOME (Node (map (fn m => Node (map (append m) elements)) elements),
model, args)
end
| _ => NONE;
fun Product_Type_fst_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹fst›, Type ("fun", [Type (\<^type_name>‹Product_Type.prod›, [T, U]), _])) =>
let
val constants_T = make_constants ctxt model T
val size_U = size_of_type ctxt model U
in
SOME (Node (maps (replicate size_U) constants_T), model, args)
end
| _ => NONE;
fun Product_Type_snd_interpreter ctxt model args t =
case t of
Const (\<^const_name>‹snd›, Type ("fun", [Type (\<^type_name>‹Product_Type.prod›, [T, U]), _])) =>
let
val size_T = size_of_type ctxt model T
val constants_U = make_constants ctxt model U
in
SOME (Node (flat (replicate size_T constants_U)), model, args)
end
| _ => NONE;
fun stlc_printer ctxt model T intr assignment =
let
val strip_leading_quote = perhaps (try (unprefix "'"))
fun string_of_typ (Type (s, _)) = s
| string_of_typ (TFree (x, _)) = strip_leading_quote x
| string_of_typ (TVar ((x,i), _)) =
strip_leading_quote x ^ string_of_int i
fun index_from_interpretation (Leaf xs) =
find_index (Prop_Logic.eval assignment) xs
| index_from_interpretation _ =
raise REFUTE ("stlc_printer",
"interpretation for ground type is not a leaf")
in
case T of
Type ("fun", [T1, T2]) =>
let
val constants = make_constants ctxt model T1
val results =
(case intr of
Node xs => xs
| _ => raise REFUTE ("stlc_printer",
"interpretation for function type is a leaf"))
val pairs = map (fn (arg, result) =>
HOLogic.mk_prod
(print ctxt model T1 arg assignment,
print ctxt model T2 result assignment))
(constants ~~ results)
val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT
val HOLogic_empty_set = Const (\<^const_abbrev>‹Set.empty›, HOLogic_setT)
val HOLogic_insert =
Const (\<^const_name>‹insert›, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
in
SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set)
end
| Type (\<^type_name>‹prop›, []) =>
(case index_from_interpretation intr of
~1 => SOME (HOLogic.mk_Trueprop (Const (\<^const_name>‹undefined›, HOLogic.boolT)))
| 0 => SOME (HOLogic.mk_Trueprop \<^term>‹True›)
| 1 => SOME (HOLogic.mk_Trueprop \<^term>‹False›)
| _ => raise REFUTE ("stlc_interpreter",
"illegal interpretation for a propositional value"))
| Type _ =>
if index_from_interpretation intr = (~1) then
SOME (Const (\<^const_name>‹undefined›, T))
else
SOME (Const (string_of_typ T ^
string_of_int (index_from_interpretation intr), T))
| TFree _ =>
if index_from_interpretation intr = (~1) then
SOME (Const (\<^const_name>‹undefined›, T))
else
SOME (Const (string_of_typ T ^
string_of_int (index_from_interpretation intr), T))
| TVar _ =>
if index_from_interpretation intr = (~1) then
SOME (Const (\<^const_name>‹undefined›, T))
else
SOME (Const (string_of_typ T ^
string_of_int (index_from_interpretation intr), T))
end;
fun set_printer ctxt model T intr assignment =
(case T of
Type (\<^type_name>‹set›, [T1]) =>
let
val constants = make_constants ctxt model T1
val results = (case intr of
Node xs => xs
| _ => raise REFUTE ("set_printer",
"interpretation for set type is a leaf"))
val elements = map_filter (fn (arg, result) =>
case result of
Leaf [fmTrue, _] =>
if Prop_Logic.eval assignment fmTrue then
SOME (print ctxt model T1 arg assignment)
else
NONE
| _ =>
raise REFUTE ("set_printer",
"illegal interpretation for a Boolean value"))
(constants ~~ results)
val HOLogic_setT1 = HOLogic.mk_setT T1
val HOLogic_empty_set = Const (\<^const_abbrev>‹Set.empty›, HOLogic_setT1)
val HOLogic_insert =
Const (\<^const_name>‹insert›, T1 --> HOLogic_setT1 --> HOLogic_setT1)
in
SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc)
(HOLogic_empty_set, elements))
end
| _ =>
NONE);
fun IDT_printer ctxt model T intr assignment =
let
val thy = Proof_Context.theory_of ctxt
in
(case T of
Type (s, Ts) =>
(case BNF_LFP_Compat.get_info thy [] s of
SOME info =>
let
val (typs, _) = model
val index = #index info
val descr = #descr info
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
val typ_assoc = dtyps ~~ Ts
val _ =
if Library.exists (fn d =>
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_printer", "datatype argument (for type " ^
Syntax.string_of_typ ctxt (Type (s, Ts)) ^ ") is not a variable")
else ()
val element =
(case intr of
Leaf xs => find_index (Prop_Logic.eval assignment) xs
| Node _ => raise REFUTE ("IDT_printer",
"interpretation is not a leaf"))
in
if element < 0 then
SOME (Const (\<^const_name>‹undefined›, Type (s, Ts)))
else
let
fun get_constr_args (cname, cargs) =
let
val cTerm = Const (cname,
map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts))
val (iC, _, _) = interpret ctxt (typs, []) {maxvars=0,
def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm
fun get_args (Leaf xs) =
if find_index (fn x => x = True) xs = element then
SOME []
else
NONE
| get_args (Node xs) =
let
fun search ([], _) =
NONE
| search (x::xs, n) =
(case get_args x of
SOME result => SOME (n::result)
| NONE => search (xs, n+1))
in
search (xs, 0)
end
in
Option.map (fn args => (cTerm, cargs, args)) (get_args iC)
end
val (cTerm, cargs, args) =
(case get_first get_constr_args constrs of
SOME x => x
| NONE => raise REFUTE ("IDT_printer",
"no matching constructor found for element " ^
string_of_int element))
val argsTerms = map (fn (d, n) =>
let
val dT = typ_of_dtyp descr typ_assoc d
val consts = make_constants ctxt (typs, []) dT
in
print ctxt (typs, []) dT (nth consts n) assignment
end) (cargs ~~ args)
in
SOME (list_comb (cTerm, argsTerms))
end
end
| NONE =>
NONE)
| _ =>
NONE)
end;
val _ =
Theory.setup
(add_interpreter "stlc" stlc_interpreter #>
add_interpreter "Pure" Pure_interpreter #>
add_interpreter "HOLogic" HOLogic_interpreter #>
add_interpreter "set" set_interpreter #>
add_interpreter "IDT" IDT_interpreter #>
add_interpreter "IDT_constructor" IDT_constructor_interpreter #>
add_interpreter "IDT_recursion" IDT_recursion_interpreter #>
add_interpreter "Finite_Set.card" Finite_Set_card_interpreter #>
add_interpreter "Finite_Set.finite" Finite_Set_finite_interpreter #>
add_interpreter "Nat_Orderings.less" Nat_less_interpreter #>
add_interpreter "Nat_HOL.plus" Nat_plus_interpreter #>
add_interpreter "Nat_HOL.minus" Nat_minus_interpreter #>
add_interpreter "Nat_HOL.times" Nat_times_interpreter #>
add_interpreter "List.append" List_append_interpreter #>
add_interpreter "Product_Type.prod.fst" Product_Type_fst_interpreter #>
add_interpreter "Product_Type.prod.snd" Product_Type_snd_interpreter #>
add_printer "stlc" stlc_printer #>
add_printer "set" set_printer #>
add_printer "IDT" IDT_printer);
val scan_parm = Parse.name -- (Scan.optional (\<^keyword>‹=› |-- Parse.name) "true")
val scan_parms = Scan.optional (\<^keyword>‹[› |-- Parse.list scan_parm --| \<^keyword>‹]›) [];
val _ =
Outer_Syntax.command \<^command_keyword>‹refute›
"try to find a model that refutes a given subgoal"
(scan_parms -- Scan.optional Parse.nat 1 >>
(fn (parms, i) =>
Toplevel.keep_proof (fn state =>
let
val ctxt = Toplevel.context_of state;
val {goal = st, ...} = Proof.raw_goal (Toplevel.proof_of state);
in refute_goal ctxt parms st i; () end)));
val _ =
Outer_Syntax.command \<^command_keyword>‹refute_params›
"show/store default parameters for the 'refute' command"
(scan_parms >> (fn parms =>
Toplevel.theory (fn thy =>
let
val thy' = fold set_default_param parms thy;
val output =
(case get_default_params (Proof_Context.init_global thy') of
[] => "none"
| new_defaults => cat_lines (map (fn (x, y) => x ^ "=" ^ y) new_defaults));
val _ = writeln ("Default parameters for 'refute':\n" ^ output);
in thy' end)));
end;