--- a/src/HOL/SPARK/Tools/spark_vcs.ML Mon Apr 18 15:02:50 2011 +0200
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML Mon Apr 18 16:33:45 2011 +0200
@@ -8,7 +8,7 @@
signature SPARK_VCS =
sig
val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules ->
- (string * string) * Fdl_Parser.vcs -> Path.T -> theory -> theory
+ (string * string) * Fdl_Parser.vcs -> Path.T -> string -> theory -> theory
val add_proof_fun: (typ option -> 'a -> term) ->
string * ((string list * string) option * 'a) ->
theory -> theory
@@ -54,7 +54,8 @@
proving: bool,
vcs: (string * thm list option *
(string * expr) list * (string * expr) list) VCtab.table,
- path: Path.T} option}
+ path: Path.T,
+ prefix: string} option}
val empty : T = {pfuns = Symtab.empty, type_map = Symtab.empty, env = NONE}
val extend = I
fun merge ({pfuns = pfuns1, type_map = type_map1, env = NONE},
@@ -72,6 +73,18 @@
val lcase_eq = (op =) o pairself (to_lower o Long_Name.base_name);
+fun lookup_prfx "" tab s = Symtab.lookup tab s
+ | lookup_prfx prfx tab s = (case Symtab.lookup tab s of
+ NONE => Symtab.lookup tab (prfx ^ "__" ^ s)
+ | x => x);
+
+fun strip_prfx s =
+ let
+ fun strip ys [] = ("", implode ys)
+ | strip ys ("_" :: "_" :: xs) = (implode (rev xs), implode ys)
+ | strip ys (x :: xs) = strip (x :: ys) xs
+ in strip [] (rev (raw_explode s)) end;
+
fun mk_unop s t =
let val T = fastype_of t
in Const (s, T --> T) $ t end;
@@ -86,14 +99,14 @@
HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u)
end;
-fun get_type thy ty =
+fun get_type thy prfx ty =
let val {type_map, ...} = VCs.get thy
- in Symtab.lookup type_map ty end;
+ in lookup_prfx prfx type_map ty end;
-fun mk_type _ "integer" = HOLogic.intT
- | mk_type _ "boolean" = HOLogic.boolT
- | mk_type thy ty =
- (case get_type thy ty of
+fun mk_type _ _ "integer" = HOLogic.intT
+ | mk_type _ _ "boolean" = HOLogic.boolT
+ | mk_type thy prfx ty =
+ (case get_type thy prfx ty of
NONE =>
Syntax.check_typ (Proof_Context.init_global thy)
(Type (Sign.full_name thy (Binding.name ty), []))
@@ -120,9 +133,9 @@
val mangle_name = strip_underscores #> strip_tilde;
-fun mk_variables thy xs ty (tab, ctxt) =
+fun mk_variables thy prfx xs ty (tab, ctxt) =
let
- val T = mk_type thy ty;
+ val T = mk_type thy prfx ty;
val (ys, ctxt') = Name.variants (map mangle_name xs) ctxt;
val zs = map (Free o rpair T) ys;
in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end;
@@ -266,7 +279,7 @@
end;
-fun check_no_assoc thy s = case get_type thy s of
+fun check_no_assoc thy prfx s = case get_type thy prfx s of
NONE => ()
| SOME _ => error ("Cannot associate a type with " ^ s ^
"\nsince it is no record or enumeration type");
@@ -280,15 +293,15 @@
else SOME ("either has no element " ^ el ^
" or it is at the wrong position");
-fun add_type_def (s, Basic_Type ty) (ids, thy) =
- (check_no_assoc thy s;
+fun add_type_def prfx (s, Basic_Type ty) (ids, thy) =
+ (check_no_assoc thy prfx s;
(ids,
Typedecl.abbrev_global (Binding.name s, [], NoSyn)
- (mk_type thy ty) thy |> snd))
+ (mk_type thy prfx ty) thy |> snd))
- | add_type_def (s, Enum_Type els) ((tab, ctxt), thy) =
+ | add_type_def prfx (s, Enum_Type els) ((tab, ctxt), thy) =
let
- val (thy', tyname) = (case get_type thy s of
+ val (thy', tyname) = (case get_type thy prfx s of
NONE =>
let
val tyb = Binding.name s;
@@ -304,9 +317,18 @@
| SOME (T as Type (tyname, [])) =>
(case Datatype_Data.get_constrs thy tyname of
NONE => assoc_ty_err thy T s "is not a datatype"
- | SOME cs => (case check_enum els cs of
- NONE => (thy, tyname)
- | SOME msg => assoc_ty_err thy T s msg)));
+ | SOME cs =>
+ let
+ val (prfx', _) = strip_prfx s;
+ val els' =
+ if prfx' = "" then els
+ else map (unprefix (prfx' ^ "__")) els
+ handle Fail _ => error ("Bad enumeration type " ^ s)
+ in
+ case check_enum els' cs of
+ NONE => (thy, tyname)
+ | SOME msg => assoc_ty_err thy T s msg
+ end));
val cs = map Const (the (Datatype_Data.get_constrs thy' tyname))
in
((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab,
@@ -314,18 +336,18 @@
thy')
end
- | add_type_def (s, Array_Type (argtys, resty)) (ids, thy) =
- (check_no_assoc thy s;
+ | add_type_def prfx (s, Array_Type (argtys, resty)) (ids, thy) =
+ (check_no_assoc thy prfx s;
(ids,
Typedecl.abbrev_global (Binding.name s, [], NoSyn)
- (foldr1 HOLogic.mk_prodT (map (mk_type thy) argtys) -->
- mk_type thy resty) thy |> snd))
+ (foldr1 HOLogic.mk_prodT (map (mk_type thy prfx) argtys) -->
+ mk_type thy prfx resty) thy |> snd))
- | add_type_def (s, Record_Type fldtys) (ids, thy) =
+ | add_type_def prfx (s, Record_Type fldtys) (ids, thy) =
(ids,
let val fldTs = maps (fn (flds, ty) =>
- map (rpair (mk_type thy ty)) flds) fldtys
- in case get_type thy s of
+ map (rpair (mk_type thy prfx ty)) flds) fldtys
+ in case get_type thy prfx s of
NONE =>
Record.add_record true ([], Binding.name s) NONE
(map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy
@@ -349,12 +371,12 @@
thy))
end)
- | add_type_def (s, Pending_Type) (ids, thy) =
- (check_no_assoc thy s;
+ | add_type_def prfx (s, Pending_Type) (ids, thy) =
+ (check_no_assoc thy prfx s;
(ids, Typedecl.typedecl_global (Binding.name s, [], NoSyn) thy |> snd));
-fun term_of_expr thy types funs pfuns =
+fun term_of_expr thy prfx types pfuns =
let
fun tm_of vs (Funct ("->", [e, e'])) =
(HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
@@ -424,7 +446,7 @@
| tm_of vs (Quantifier (s, xs, ty, e)) =
let
- val (ys, vs') = mk_variables thy xs ty vs;
+ val (ys, vs') = mk_variables thy prfx xs ty vs;
val q = (case s of
"for_all" => HOLogic.mk_all
| "for_some" => HOLogic.mk_exists)
@@ -442,7 +464,7 @@
[e] =>
let
val (t, rcdty) = tm_of vs e;
- val rT = mk_type thy rcdty
+ val rT = mk_type thy prfx rcdty
in case (get_record_info thy rT, lookup types rcdty) of
(SOME {fields, ...}, SOME (Record_Type fldtys)) =>
(case (find_field fname fields,
@@ -462,9 +484,9 @@
[e, e'] =>
let
val (t, rcdty) = tm_of vs e;
- val rT = mk_type thy rcdty;
+ val rT = mk_type thy prfx rcdty;
val (u, fldty) = tm_of vs e';
- val fT = mk_type thy fldty
+ val fT = mk_type thy prfx fldty
in case get_record_info thy rT of
SOME {fields, ...} =>
(case find_field fname fields of
@@ -490,7 +512,8 @@
(case try (unsuffix "__pos") s of
SOME tyname => (case es of
[e] => (Const (@{const_name pos},
- mk_type thy tyname --> HOLogic.intT) $ fst (tm_of vs e), integerN)
+ mk_type thy prfx tyname --> HOLogic.intT) $ fst (tm_of vs e),
+ integerN)
| _ => error ("Function " ^ s ^ " expects one argument"))
| NONE =>
@@ -498,7 +521,8 @@
(case try (unsuffix "__val") s of
SOME tyname => (case es of
[e] => (Const (@{const_name val},
- HOLogic.intT --> mk_type thy tyname) $ fst (tm_of vs e), tyname)
+ HOLogic.intT --> mk_type thy prfx tyname) $ fst (tm_of vs e),
+ tyname)
| _ => error ("Function " ^ s ^ " expects one argument"))
| NONE =>
@@ -507,7 +531,7 @@
[e] =>
let
val (t, tyname) = tm_of vs e;
- val T = mk_type thy tyname
+ val T = mk_type thy prfx tyname
in (Const
(if s = "succ" then @{const_name succ}
else @{const_name pred}, T --> T) $ t, tyname)
@@ -516,7 +540,7 @@
(* user-defined proof function *)
else
- (case Symtab.lookup pfuns s of
+ (case lookup_prfx prfx pfuns s of
SOME (SOME (_, resty), t) =>
(list_comb (t, map (fst o tm_of vs) es), resty)
| _ => error ("Undeclared proof function " ^ s))))))
@@ -534,8 +558,9 @@
in case lookup types ty of
SOME (Array_Type (idxtys, elty)) =>
let
- val T = foldr1 HOLogic.mk_prodT (map (mk_type thy) idxtys);
- val U = mk_type thy elty;
+ val T = foldr1 HOLogic.mk_prodT
+ (map (mk_type thy prfx) idxtys);
+ val U = mk_type thy prfx elty;
val fT = T --> U
in
(Const (@{const_name fun_upd}, fT --> T --> U --> fT) $
@@ -548,7 +573,7 @@
| tm_of vs (Record (s, flds)) =
let
- val T = mk_type thy s;
+ val T = mk_type thy prfx s;
val {extension = (ext_name, _), fields, ...} =
(case get_record_info thy T of
NONE => error (s ^ " is not a record type")
@@ -572,7 +597,7 @@
in
(list_comb
(Const (ext_name,
- map (mk_type thy) ftys @ [HOLogic.unitT] ---> T),
+ map (mk_type thy prfx) ftys @ [HOLogic.unitT] ---> T),
fvals @ [HOLogic.unit]),
s)
end
@@ -581,9 +606,9 @@
(case lookup types s of
SOME (Array_Type (idxtys, elty)) =>
let
- val Ts = map (mk_type thy) idxtys;
+ val Ts = map (mk_type thy prfx) idxtys;
val T = foldr1 HOLogic.mk_prodT Ts;
- val U = mk_type thy elty;
+ val U = mk_type thy prfx elty;
fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)]
| mk_idx' T (e, SOME e') = Const (@{const_name atLeastAtMost},
T --> T --> HOLogic.mk_setT T) $
@@ -618,8 +643,8 @@
in tm_of end;
-fun term_of_rule thy types funs pfuns ids rule =
- let val tm_of = fst o term_of_expr thy types funs pfuns ids
+fun term_of_rule thy prfx types pfuns ids rule =
+ let val tm_of = fst o term_of_expr thy prfx types pfuns ids
in case rule of
Inference_Rule (es, e) => Logic.list_implies
(map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e))
@@ -676,14 +701,14 @@
val add_expr_idents = fold_expr (K I) (insert (op =));
-fun pfun_type thy (argtys, resty) =
- map (mk_type thy) argtys ---> mk_type thy resty;
+fun pfun_type thy prfx (argtys, resty) =
+ map (mk_type thy prfx) argtys ---> mk_type thy prfx resty;
-fun check_pfun_type thy s t optty1 optty2 =
+fun check_pfun_type thy prfx s t optty1 optty2 =
let
val T = fastype_of t;
fun check ty =
- let val U = pfun_type thy ty
+ let val U = pfun_type thy prfx ty
in
T = U orelse
error ("Type\n" ^
@@ -698,11 +723,13 @@
fun upd_option x y = if is_some x then x else y;
-fun check_pfuns_types thy funs =
+fun check_pfuns_types thy prfx funs =
Symtab.map (fn s => fn (optty, t) =>
- let val optty' = lookup funs s
+ let val optty' = lookup funs
+ (if prfx = "" then s
+ else unprefix (prfx ^ "__") s handle Fail _ => s)
in
- (check_pfun_type thy s t optty optty';
+ (check_pfun_type thy prfx s t optty optty';
(NONE |> upd_option optty |> upd_option optty', t))
end);
@@ -713,9 +740,9 @@
(Pretty.big_list "The following verification conditions have not been proved:"
(map Pretty.str names)))
-fun set_env (env as {funs, ...}) thy = VCs.map (fn
+fun set_env (env as {funs, prefix, ...}) thy = VCs.map (fn
{pfuns, type_map, env = NONE} =>
- {pfuns = check_pfuns_types thy funs pfuns,
+ {pfuns = check_pfuns_types thy prefix funs pfuns,
type_map = type_map,
env = SOME env}
| _ => err_unfinished ()) thy;
@@ -724,9 +751,9 @@
SOME i => [HOLogic.mk_Trueprop (Var (("C", i), HOLogic.boolT))]
| NONE => error ("Bad conclusion identifier: C" ^ s));
-fun mk_vc thy types funs pfuns ids (tr, proved, ps, cs) =
+fun mk_vc thy prfx types pfuns ids (tr, proved, ps, cs) =
let val prop_of =
- HOLogic.mk_Trueprop o fst o term_of_expr thy types funs pfuns ids
+ HOLogic.mk_Trueprop o fst o term_of_expr thy prfx types pfuns ids
in
(tr, proved,
Element.Assumes (map (fn (s', e) =>
@@ -738,16 +765,16 @@
fun fold_vcs f vcs =
VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs;
-fun pfuns_of_vcs pfuns vcs =
+fun pfuns_of_vcs prfx pfuns vcs =
fold_vcs (add_expr_pfuns o snd) vcs [] |>
- filter_out (Symtab.defined pfuns);
+ filter (is_none o lookup_prfx prfx pfuns);
-fun declare_missing_pfuns thy funs pfuns vcs (tab, ctxt) =
+fun declare_missing_pfuns thy prfx funs pfuns vcs (tab, ctxt) =
let
val (fs, (tys, Ts)) =
- pfuns_of_vcs pfuns vcs |>
+ pfuns_of_vcs prfx pfuns vcs |>
map_filter (fn s => lookup funs s |>
- Option.map (fn ty => (s, (SOME ty, pfun_type thy ty)))) |>
+ Option.map (fn ty => (s, (SOME ty, pfun_type thy prfx ty)))) |>
split_list ||> split_list;
val (fs', ctxt') = Name.variants fs ctxt
in
@@ -762,11 +789,11 @@
{env = SOME {proving = true, ...}, ...} => err_unfinished ()
| {pfuns, type_map, env} =>
let
- val optty' = (case env of
- SOME {funs, ...} => lookup funs s
- | NONE => NONE);
+ val (optty', prfx) = (case env of
+ SOME {funs, prefix, ...} => (lookup funs s, prefix)
+ | NONE => (NONE, ""));
val optty'' = NONE |> upd_option optty |> upd_option optty';
- val t = prep (Option.map (pfun_type thy) optty'') raw_t;
+ val t = prep (Option.map (pfun_type thy prfx) optty'') raw_t;
val _ = (case fold_aterms (fn u =>
if is_Var u orelse is_Free u then insert (op =) u else I) t [] of
[] => ()
@@ -775,7 +802,7 @@
" contains free variable(s) " ^
commas (map (Syntax.string_of_term_global thy) ts)));
in
- (check_pfun_type thy s t optty optty';
+ (check_pfun_type thy prfx s t optty optty';
if is_some optty'' orelse is_none env then
{pfuns = Symtab.update_new (s, (optty'', t)) pfuns,
type_map = type_map,
@@ -813,29 +840,29 @@
fun lookup_vc thy name =
(case VCs.get thy of
- {env = SOME {vcs, types, funs, ids, ctxt, ...}, pfuns, ...} =>
+ {env = SOME {vcs, types, funs, ids, ctxt, prefix, ...}, pfuns, ...} =>
(case VCtab.lookup vcs name of
- SOME vc =>
+ SOME vc =>
let val (pfuns', ctxt', ids') =
- declare_missing_pfuns thy funs pfuns vcs ids
- in SOME (ctxt @ [ctxt'], mk_vc thy types funs pfuns' ids' vc) end
+ declare_missing_pfuns thy prefix funs pfuns vcs ids
+ in SOME (ctxt @ [ctxt'], mk_vc thy prefix types pfuns' ids' vc) end
| NONE => NONE)
| _ => NONE);
fun get_vcs thy = (case VCs.get thy of
- {env = SOME {vcs, types, funs, ids, ctxt, defs, ...}, pfuns, ...} =>
+ {env = SOME {vcs, types, funs, ids, ctxt, defs, prefix, ...}, pfuns, ...} =>
let val (pfuns', ctxt', ids') =
- declare_missing_pfuns thy funs pfuns vcs ids
+ declare_missing_pfuns thy prefix funs pfuns vcs ids
in
(ctxt @ [ctxt'], defs,
VCtab.dest vcs |>
- map (apsnd (mk_vc thy types funs pfuns' ids')))
+ map (apsnd (mk_vc thy prefix types pfuns' ids')))
end
| _ => ([], [], []));
fun mark_proved name thms = VCs.map (fn
{pfuns, type_map,
- env = SOME {ctxt, defs, types, funs, ids, vcs, path, ...}} =>
+ env = SOME {ctxt, defs, types, funs, ids, vcs, path, prefix, ...}} =>
{pfuns = pfuns,
type_map = type_map,
env = SOME {ctxt = ctxt, defs = defs,
@@ -843,7 +870,8 @@
proving = true,
vcs = VCtab.map_entry name (fn (trace, _, ps, cs) =>
(trace, SOME thms, ps, cs)) vcs,
- path = path}}
+ path = path,
+ prefix = prefix}}
| x => x);
fun close thy =
@@ -878,9 +906,9 @@
fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i);
-fun add_const (s, ty) ((tab, ctxt), thy) =
+fun add_const prfx (s, ty) ((tab, ctxt), thy) =
let
- val T = mk_type thy ty;
+ val T = mk_type thy prfx ty;
val b = Binding.name s;
val c = Const (Sign.full_name thy b, T)
in
@@ -889,13 +917,13 @@
Sign.add_consts_i [(b, T, NoSyn)] thy))
end;
-fun add_def types funs pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) =
+fun add_def prfx types pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) =
(case lookup consts s of
SOME ty =>
let
- val (t, ty') = term_of_expr thy types funs pfuns ids e;
- val T = mk_type thy ty;
- val T' = mk_type thy ty';
+ val (t, ty') = term_of_expr thy prfx types pfuns ids e;
+ val T = mk_type thy prfx ty;
+ val T' = mk_type thy prfx ty';
val _ = T = T' orelse
error ("Declared type " ^ ty ^ " of " ^ s ^
"\ndoes not match type " ^ ty' ^ " in definition");
@@ -913,12 +941,12 @@
end
| NONE => error ("Undeclared constant " ^ s));
-fun add_var (s, ty) (ids, thy) =
- let val ([Free p], ids') = mk_variables thy [s] ty ids
+fun add_var prfx (s, ty) (ids, thy) =
+ let val ([Free p], ids') = mk_variables thy prfx [s] ty ids
in (p, (ids', thy)) end;
-fun add_init_vars vcs (ids_thy as ((tab, _), _)) =
- fold_map add_var
+fun add_init_vars prfx vcs (ids_thy as ((tab, _), _)) =
+ fold_map (add_var prfx)
(map_filter
(fn s => case try (unsuffix "~") s of
SOME s' => (case Symtab.lookup tab s' of
@@ -935,20 +963,20 @@
(map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules);
(* sort definitions according to their dependency *)
-fun sort_defs _ _ [] sdefs = rev sdefs
- | sort_defs pfuns consts defs sdefs =
+fun sort_defs _ _ _ [] sdefs = rev sdefs
+ | sort_defs prfx pfuns consts defs sdefs =
(case find_first (fn (_, (_, e)) =>
- forall (Symtab.defined pfuns) (add_expr_pfuns e []) andalso
+ forall (is_some o lookup_prfx prfx pfuns) (add_expr_pfuns e []) andalso
forall (fn id =>
member (fn (s, (_, (s', _))) => s = s') sdefs id orelse
consts id)
(add_expr_idents e [])) defs of
- SOME d => sort_defs pfuns consts
+ SOME d => sort_defs prfx pfuns consts
(remove (op =) d defs) (d :: sdefs)
| NONE => error ("Bad definitions: " ^ rulenames defs));
fun set_vcs ({types, vars, consts, funs} : decls)
- (rules, _) ((_, ident), vcs) path thy =
+ (rules, _) ((_, ident), vcs) path prfx thy =
let
val {pfuns, ...} = VCs.get thy;
val (defs, rules') = partition_opt dest_def rules;
@@ -965,7 +993,7 @@
(filter_out (is_trivial_vc o snd) vcs)) vcs);
val _ = (case filter_out (is_some o lookup funs)
- (pfuns_of_vcs pfuns vcs') of
+ (pfuns_of_vcs prfx pfuns vcs') of
[] => ()
| fs => error ("Undeclared proof function(s) " ^ commas fs));
@@ -975,27 +1003,27 @@
Symtab.update ("true", (HOLogic.true_const, booleanN)),
Name.context),
thy |> Sign.add_path (Long_Name.base_name ident)) |>
- fold add_type_def (items types) |>
- fold (snd oo add_const) consts' |> (fn ids_thy as ((tab, _), _) =>
+ fold (add_type_def prfx) (items types) |>
+ fold (snd oo add_const prfx) consts' |> (fn ids_thy as ((tab, _), _) =>
ids_thy |>
- fold_map (add_def types funs pfuns consts)
- (sort_defs pfuns (Symtab.defined tab) defs []) ||>>
- fold_map add_var (items vars) ||>>
- add_init_vars vcs');
+ fold_map (add_def prfx types pfuns consts)
+ (sort_defs prfx pfuns (Symtab.defined tab) defs []) ||>>
+ fold_map (add_var prfx) (items vars) ||>>
+ add_init_vars prfx vcs');
val ctxt =
[Element.Fixes (map (fn (s, T) =>
(Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)),
Element.Assumes (map (fn (id, rl) =>
((mk_rulename id, []),
- [(term_of_rule thy' types funs pfuns ids rl, [])]))
+ [(term_of_rule thy' prfx types pfuns ids rl, [])]))
other_rules),
Element.Notes (Thm.definitionK,
[((Binding.name "defns", []), map (rpair [] o single o snd) defs')])]
in
set_env {ctxt = ctxt, defs = defs', types = types, funs = funs,
- ids = ids, proving = false, vcs = vcs', path = path} thy'
+ ids = ids, proving = false, vcs = vcs', path = path, prefix = prfx} thy'
end;
end;