--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,870 @@
+(* Title: HOL/SPARK/Tools/spark_vcs.ML
+ Author: Stefan Berghofer
+ Copyright: secunet Security Networks AG
+
+Store for verification conditions generated by SPARK/Ada.
+*)
+
+signature SPARK_VCS =
+sig
+ val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules -> Fdl_Parser.vcs ->
+ Path.T -> theory -> theory
+ val add_proof_fun: (typ option -> 'a -> term) ->
+ string * ((string list * string) option * 'a) ->
+ theory -> theory
+ val lookup_vc: theory -> string -> (Element.context_i list *
+ (string * bool * Element.context_i * Element.statement_i)) option
+ val get_vcs: theory ->
+ Element.context_i list * (binding * thm) list *
+ (string * (string * bool * Element.context_i * Element.statement_i)) list
+ val mark_proved: string -> theory -> theory
+ val close: theory -> theory
+ val is_closed: theory -> bool
+end;
+
+structure SPARK_VCs: SPARK_VCS =
+struct
+
+open Fdl_Parser;
+
+
+(** utilities **)
+
+fun mk_unop s t =
+ let val T = fastype_of t
+ in Const (s, T --> T) $ t end;
+
+fun mk_times (t, u) =
+ let
+ val setT = fastype_of t;
+ val T = HOLogic.dest_setT setT;
+ val U = HOLogic.dest_setT (fastype_of u)
+ in
+ Const (@{const_name Sigma}, setT --> (T --> HOLogic.mk_setT U) -->
+ HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u)
+ end;
+
+fun mk_type _ "integer" = HOLogic.intT
+ | mk_type _ "boolean" = HOLogic.boolT
+ | mk_type thy ty = Syntax.check_typ (ProofContext.init_global thy)
+ (Type (Sign.full_name thy (Binding.name ty), []));
+
+val booleanN = "boolean";
+val integerN = "integer";
+
+fun mk_qual_name thy s s' =
+ Sign.full_name thy (Binding.qualify true s (Binding.name s'));
+
+fun define_overloaded (def_name, eq) lthy =
+ let
+ val ((c, _), rhs) = eq |> Syntax.check_term lthy |>
+ Logic.dest_equals |>> dest_Free;
+ val ((_, (_, thm)), lthy') = Local_Theory.define
+ ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy
+ val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy');
+ val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm
+ in (thm', lthy') end;
+
+fun strip_underscores s =
+ strip_underscores (unsuffix "_" s) handle Fail _ => s;
+
+fun strip_tilde s =
+ unsuffix "~" s ^ "_init" handle Fail _ => s;
+
+val mangle_name = strip_underscores #> strip_tilde;
+
+fun mk_variables thy xs ty (tab, ctxt) =
+ let
+ val T = mk_type thy 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;
+
+
+(** generate properties of enumeration types **)
+
+fun add_enum_type tyname els (tab, ctxt) thy =
+ let
+ val tyb = Binding.name tyname;
+ val tyname' = Sign.full_name thy tyb;
+ val T = Type (tyname', []);
+ val case_name = mk_qual_name thy tyname (tyname ^ "_case");
+ val cs = map (fn s => Const (mk_qual_name thy tyname s, T)) els;
+ val k = length els;
+ val p = Const (@{const_name pos}, T --> HOLogic.intT);
+ val v = Const (@{const_name val}, HOLogic.intT --> T);
+ val card = Const (@{const_name card},
+ HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T;
+
+ fun mk_binrel_def s f = Logic.mk_equals
+ (Const (s, T --> T --> HOLogic.boolT),
+ Abs ("x", T, Abs ("y", T,
+ Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $
+ (f $ Bound 1) $ (f $ Bound 0))));
+
+ val (((def1, def2), def3), lthy) = thy |>
+ Datatype.add_datatype {strict = true, quiet = true} [tyname]
+ [([], tyb, NoSyn,
+ map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
+
+ Class.instantiation ([tyname'], [], @{sort enum}) |>
+
+ define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals
+ (p,
+ list_comb (Const (case_name, replicate k HOLogic.intT @
+ [T] ---> HOLogic.intT),
+ map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>>
+
+ define_overloaded ("less_eq_" ^ tyname ^ "_def",
+ mk_binrel_def @{const_name less_eq} p) ||>>
+ define_overloaded ("less_" ^ tyname ^ "_def",
+ mk_binrel_def @{const_name less} p);
+
+ val UNIV_eq = Goal.prove lthy [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_UNIV T, HOLogic.mk_set T cs)))
+ (fn _ =>
+ rtac @{thm subset_antisym} 1 THEN
+ rtac @{thm subsetI} 1 THEN
+ Datatype_Aux.exh_tac (K (#exhaust (Datatype_Data.the_info
+ (ProofContext.theory_of lthy) tyname'))) 1 THEN
+ ALLGOALS (asm_full_simp_tac (simpset_of lthy)));
+
+ val finite_UNIV = Goal.prove lthy [] []
+ (HOLogic.mk_Trueprop (Const (@{const_name finite},
+ HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T))
+ (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1);
+
+ val card_UNIV = Goal.prove lthy [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (card, HOLogic.mk_number HOLogic.natT k)))
+ (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1);
+
+ val range_pos = Goal.prove lthy [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (@{const_name image}, (T --> HOLogic.intT) -->
+ HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $
+ p $ HOLogic.mk_UNIV T,
+ Const (@{const_name atLeastLessThan}, HOLogic.intT -->
+ HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $
+ HOLogic.mk_number HOLogic.intT 0 $
+ (@{term int} $ card))))
+ (fn _ =>
+ simp_tac (simpset_of lthy addsimps [card_UNIV]) 1 THEN
+ simp_tac (simpset_of lthy addsimps [UNIV_eq, def1]) 1 THEN
+ rtac @{thm subset_antisym} 1 THEN
+ simp_tac (simpset_of lthy) 1 THEN
+ rtac @{thm subsetI} 1 THEN
+ asm_full_simp_tac (simpset_of lthy addsimps @{thms interval_expand}
+ delsimps @{thms atLeastLessThan_iff}) 1);
+
+ val lthy' =
+ Class.prove_instantiation_instance (fn _ =>
+ Class.intro_classes_tac [] THEN
+ rtac finite_UNIV 1 THEN
+ rtac range_pos 1 THEN
+ simp_tac (HOL_basic_ss addsimps [def3]) 1 THEN
+ simp_tac (HOL_basic_ss addsimps [def2]) 1) lthy;
+
+ val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) =>
+ let
+ val n = HOLogic.mk_number HOLogic.intT i;
+ val th = Goal.prove lthy' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n)))
+ (fn _ => simp_tac (simpset_of lthy' addsimps [def1]) 1);
+ val th' = Goal.prove lthy' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c)))
+ (fn _ =>
+ rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN
+ simp_tac (simpset_of lthy' addsimps
+ [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
+ in (th, th') end) cs);
+
+ val first_el = Goal.prove lthy' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (@{const_name first_el}, T), hd cs)))
+ (fn _ => simp_tac (simpset_of lthy' addsimps
+ [@{thm first_el_def}, hd val_eqs]) 1);
+
+ val last_el = Goal.prove lthy' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (@{const_name last_el}, T), List.last cs)))
+ (fn _ => simp_tac (simpset_of lthy' addsimps
+ [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
+
+ val simp_att = [Attrib.internal (K Simplifier.simp_add)]
+
+ in
+ ((fold (Symtab.update_new o apsnd (rpair tyname)) (els ~~ cs) tab,
+ fold Name.declare els ctxt),
+ lthy' |>
+ Local_Theory.note
+ ((Binding.name (tyname ^ "_card_UNIV"), simp_att), [card_UNIV]) ||>>
+ Local_Theory.note
+ ((Binding.name (tyname ^ "_pos"), simp_att), pos_eqs) ||>>
+ Local_Theory.note
+ ((Binding.name (tyname ^ "_val"), simp_att), val_eqs) ||>>
+ Local_Theory.note
+ ((Binding.name (tyname ^ "_first_el"), simp_att), [first_el]) ||>>
+ Local_Theory.note
+ ((Binding.name (tyname ^ "_last_el"), simp_att), [last_el]) |> snd |>
+ Local_Theory.exit_global)
+ end;
+
+
+fun add_type_def (s, Basic_Type ty) (ids, thy) =
+ (ids,
+ Typedecl.abbrev_global (Binding.name s, [], NoSyn)
+ (mk_type thy ty) thy |> snd)
+
+ | add_type_def (s, Enum_Type els) (ids, thy) = add_enum_type s els ids thy
+
+ | add_type_def (s, Array_Type (argtys, resty)) (ids, thy) =
+ (ids,
+ Typedecl.abbrev_global (Binding.name s, [], NoSyn)
+ (foldr1 HOLogic.mk_prodT (map (mk_type thy) argtys) -->
+ mk_type thy resty) thy |> snd)
+
+ | add_type_def (s, Record_Type fldtys) (ids, thy) =
+ (ids,
+ Record.add_record true ([], Binding.name s) NONE
+ (maps (fn (flds, ty) =>
+ let val T = mk_type thy ty
+ in map (fn fld => (Binding.name fld, T, NoSyn)) flds
+ end) fldtys) thy)
+
+ | add_type_def (s, Pending_Type) (ids, thy) =
+ (ids, Typedecl.typedecl_global (Binding.name s, [], NoSyn) thy |> snd);
+
+
+fun term_of_expr thy types funs pfuns =
+ let
+ fun tm_of vs (Funct ("->", [e, e'])) =
+ (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct ("<->", [e, e'])) =
+ (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct ("or", [e, e'])) =
+ (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct ("and", [e, e'])) =
+ (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct ("not", [e])) =
+ (HOLogic.mk_not (fst (tm_of vs e)), booleanN)
+
+ | tm_of vs (Funct ("=", [e, e'])) =
+ (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct ("<>", [e, e'])) = (HOLogic.mk_not
+ (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN)
+
+ | tm_of vs (Funct ("<", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
+ (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct (">", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
+ (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
+
+ | tm_of vs (Funct ("<=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
+ (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+ | tm_of vs (Funct (">=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
+ (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
+
+ | tm_of vs (Funct ("+", [e, e'])) = (HOLogic.mk_binop @{const_name plus}
+ (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+ | tm_of vs (Funct ("-", [e, e'])) = (HOLogic.mk_binop @{const_name minus}
+ (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+ | tm_of vs (Funct ("*", [e, e'])) = (HOLogic.mk_binop @{const_name times}
+ (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+ | tm_of vs (Funct ("/", [e, e'])) = (HOLogic.mk_binop @{const_name divide}
+ (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+ | tm_of vs (Funct ("div", [e, e'])) = (HOLogic.mk_binop @{const_name sdiv}
+ (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+ | tm_of vs (Funct ("mod", [e, e'])) = (HOLogic.mk_binop @{const_name smod}
+ (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+ | tm_of vs (Funct ("-", [e])) =
+ (mk_unop @{const_name uminus} (fst (tm_of vs e)), integerN)
+
+ | tm_of vs (Funct ("**", [e, e'])) =
+ (Const (@{const_name power}, HOLogic.intT --> HOLogic.natT -->
+ HOLogic.intT) $ fst (tm_of vs e) $
+ (@{const nat} $ fst (tm_of vs e')), integerN)
+
+ | tm_of (tab, _) (Ident s) =
+ (case Symtab.lookup tab s of
+ SOME t_ty => t_ty
+ | NONE => error ("Undeclared identifier " ^ s))
+
+ | tm_of _ (Number i) = (HOLogic.mk_number HOLogic.intT i, integerN)
+
+ | tm_of vs (Quantifier (s, xs, ty, e)) =
+ let
+ val (ys, vs') = mk_variables thy xs ty vs;
+ val q = (case s of
+ "for_all" => HOLogic.mk_all
+ | "for_some" => HOLogic.mk_exists)
+ in
+ (fold_rev (fn Free (x, T) => fn t => q (x, T, t))
+ ys (fst (tm_of vs' e)),
+ booleanN)
+ end
+
+ | tm_of vs (Funct (s, es)) =
+
+ (* record field selection *)
+ (case try (unprefix "fld_") s of
+ SOME fname => (case es of
+ [e] =>
+ let val (t, rcdty) = tm_of vs e
+ in case lookup types rcdty of
+ SOME (Record_Type fldtys) =>
+ (case get_first (fn (flds, fldty) =>
+ if member (op =) flds fname then SOME fldty
+ else NONE) fldtys of
+ SOME fldty =>
+ (Const (mk_qual_name thy rcdty fname,
+ mk_type thy rcdty --> mk_type thy fldty) $ t,
+ fldty)
+ | NONE => error ("Record " ^ rcdty ^
+ " has no field named " ^ fname))
+ | _ => error (rcdty ^ " is not a record type")
+ end
+ | _ => error ("Function " ^ s ^ " expects one argument"))
+ | NONE =>
+
+ (* record field update *)
+ (case try (unprefix "upf_") s of
+ SOME fname => (case es of
+ [e, e'] =>
+ let
+ val (t, rcdty) = tm_of vs e;
+ val rT = mk_type thy rcdty;
+ val (u, fldty) = tm_of vs e';
+ val fT = mk_type thy fldty
+ in case lookup types rcdty of
+ SOME (Record_Type fldtys) =>
+ (case get_first (fn (flds, fldty) =>
+ if member (op =) flds fname then SOME fldty
+ else NONE) fldtys of
+ SOME fldty' =>
+ if fldty = fldty' then
+ (Const (mk_qual_name thy rcdty (fname ^ "_update"),
+ (fT --> fT) --> rT --> rT) $
+ Abs ("x", fT, u) $ t,
+ rcdty)
+ else error ("Type " ^ fldty ^
+ " does not match type " ^ fldty' ^ " of field " ^
+ fname)
+ | NONE => error ("Record " ^ rcdty ^
+ " has no field named " ^ fname))
+ | _ => error (rcdty ^ " is not a record type")
+ end
+ | _ => error ("Function " ^ s ^ " expects two arguments"))
+ | NONE =>
+
+ (* enumeration type to integer *)
+ (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)
+ | _ => error ("Function " ^ s ^ " expects one argument"))
+ | NONE =>
+
+ (* integer to enumeration type *)
+ (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)
+ | _ => error ("Function " ^ s ^ " expects one argument"))
+ | NONE =>
+
+ (* successor / predecessor of enumeration type element *)
+ if s = "succ" orelse s = "pred" then (case es of
+ [e] =>
+ let
+ val (t, tyname) = tm_of vs e;
+ val T = mk_type thy tyname
+ in (Const
+ (if s = "succ" then @{const_name succ}
+ else @{const_name pred}, T --> T) $ t, tyname)
+ end
+ | _ => error ("Function " ^ s ^ " expects one argument"))
+
+ (* user-defined proof function *)
+ else
+ (case Symtab.lookup pfuns s of
+ SOME (SOME (_, resty), t) =>
+ (list_comb (t, map (fst o tm_of vs) es), resty)
+ | _ => error ("Undeclared proof function " ^ s))))))
+
+ | tm_of vs (Element (e, es)) =
+ let val (t, ty) = tm_of vs e
+ in case lookup types ty of
+ SOME (Array_Type (_, elty)) =>
+ (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty)
+ | _ => error (ty ^ " is not an array type")
+ end
+
+ | tm_of vs (Update (e, es, e')) =
+ let val (t, ty) = tm_of vs e
+ 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 fT = T --> U
+ in
+ (Const (@{const_name fun_upd}, fT --> T --> U --> fT) $
+ t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $
+ fst (tm_of vs e'),
+ ty)
+ end
+ | _ => error (ty ^ " is not an array type")
+ end
+
+ | tm_of vs (Record (s, flds)) =
+ (case lookup types s of
+ SOME (Record_Type fldtys) =>
+ let
+ val flds' = map (apsnd (tm_of vs)) flds;
+ val fnames = maps fst fldtys;
+ val fnames' = map fst flds;
+ val (fvals, ftys) = split_list (map (fn s' =>
+ case AList.lookup (op =) flds' s' of
+ SOME fval_ty => fval_ty
+ | NONE => error ("Field " ^ s' ^ " missing in record " ^ s))
+ fnames);
+ val _ = (case subtract (op =) fnames fnames' of
+ [] => ()
+ | xs => error ("Extra field(s) " ^ commas xs ^
+ " in record " ^ s));
+ val _ = (case duplicates (op =) fnames' of
+ [] => ()
+ | xs => error ("Duplicate field(s) " ^ commas xs ^
+ " in record " ^ s))
+ in
+ (list_comb
+ (Const (mk_qual_name thy s (s ^ "_ext"),
+ map (mk_type thy) ftys @ [HOLogic.unitT] --->
+ mk_type thy s),
+ fvals @ [HOLogic.unit]),
+ s)
+ end
+ | _ => error (s ^ " is not a record type"))
+
+ | tm_of vs (Array (s, default, assocs)) =
+ (case lookup types s of
+ SOME (Array_Type (idxtys, elty)) =>
+ let
+ val Ts = map (mk_type thy) idxtys;
+ val T = foldr1 HOLogic.mk_prodT Ts;
+ val U = mk_type thy 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) $
+ fst (tm_of vs e) $ fst (tm_of vs e');
+ fun mk_idx idx =
+ if length Ts <> length idx then
+ error ("Arity mismatch in construction of array " ^ s)
+ else foldr1 mk_times (map2 mk_idx' Ts idx);
+ fun mk_upd (idxs, e) t =
+ if length idxs = 1 andalso forall (is_none o snd) (hd idxs)
+ then
+ Const (@{const_name fun_upd}, (T --> U) -->
+ T --> U --> T --> U) $ t $
+ foldl1 HOLogic.mk_prod
+ (map (fst o tm_of vs o fst) (hd idxs)) $
+ fst (tm_of vs e)
+ else
+ Const (@{const_name fun_upds}, (T --> U) -->
+ HOLogic.mk_setT T --> U --> T --> U) $ t $
+ foldl1 (HOLogic.mk_binop @{const_name sup})
+ (map mk_idx idxs) $
+ fst (tm_of vs e)
+ in
+ (fold mk_upd assocs
+ (case default of
+ SOME e => Abs ("x", T, fst (tm_of vs e))
+ | NONE => Const (@{const_name undefined}, T --> U)),
+ s)
+ end
+ | _ => error (s ^ " is not an array type"))
+
+ 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
+ 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))
+ | Substitution_Rule (es, e, e') => Logic.list_implies
+ (map (HOLogic.mk_Trueprop o tm_of) es,
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e')))
+ end;
+
+
+val builtin = Symtab.make (map (rpair ())
+ ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=",
+ "+", "-", "*", "/", "div", "mod", "**"]);
+
+fun complex_expr (Number _) = false
+ | complex_expr (Ident _) = false
+ | complex_expr (Funct (s, es)) =
+ not (Symtab.defined builtin s) orelse exists complex_expr es
+ | complex_expr (Quantifier (_, _, _, e)) = complex_expr e
+ | complex_expr _ = true;
+
+fun complex_rule (Inference_Rule (es, e)) =
+ complex_expr e orelse exists complex_expr es
+ | complex_rule (Substitution_Rule (es, e, e')) =
+ complex_expr e orelse complex_expr e' orelse
+ exists complex_expr es;
+
+val is_pfun =
+ Symtab.defined builtin orf
+ can (unprefix "fld_") orf can (unprefix "upf_") orf
+ can (unsuffix "__pos") orf can (unsuffix "__val") orf
+ equal "succ" orf equal "pred";
+
+fun fold_opt f = the_default I o Option.map f;
+fun fold_pair f g (x, y) = f x #> g y;
+
+fun fold_expr f g (Funct (s, es)) = f s #> fold (fold_expr f g) es
+ | fold_expr f g (Ident s) = g s
+ | fold_expr f g (Number _) = I
+ | fold_expr f g (Quantifier (_, _, _, e)) = fold_expr f g e
+ | fold_expr f g (Element (e, es)) =
+ fold_expr f g e #> fold (fold_expr f g) es
+ | fold_expr f g (Update (e, es, e')) =
+ fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e'
+ | fold_expr f g (Record (_, flds)) = fold (fold_expr f g o snd) flds
+ | fold_expr f g (Array (_, default, assocs)) =
+ fold_opt (fold_expr f g) default #>
+ fold (fold_pair
+ (fold (fold (fold_pair
+ (fold_expr f g) (fold_opt (fold_expr f g)))))
+ (fold_expr f g)) assocs;
+
+val add_expr_pfuns = fold_expr
+ (fn s => if is_pfun s then I else insert (op =) s) (K I);
+
+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 check_pfun_type thy s t optty1 optty2 =
+ let
+ val T = fastype_of t;
+ fun check ty =
+ let val U = pfun_type thy ty
+ in
+ T = U orelse
+ error ("Type\n" ^
+ Syntax.string_of_typ_global thy T ^
+ "\nof function " ^
+ Syntax.string_of_term_global thy t ^
+ " associated with proof function " ^ s ^
+ "\ndoes not match declared type\n" ^
+ Syntax.string_of_typ_global thy U)
+ end
+ in (Option.map check optty1; Option.map check optty2; ()) end;
+
+fun upd_option x y = if is_some x then x else y;
+
+fun check_pfuns_types thy funs =
+ Symtab.map (fn s => fn (optty, t) =>
+ let val optty' = lookup funs s
+ in
+ (check_pfun_type thy s t optty optty';
+ (NONE |> upd_option optty |> upd_option optty', t))
+ end);
+
+
+(** the VC store **)
+
+fun err_unfinished () = error "An unfinished SPARK environment is still open."
+
+fun err_vcs names = error (Pretty.string_of
+ (Pretty.big_list "The following verification conditions have not been proved:"
+ (map Pretty.str names)))
+
+val strip_number = pairself implode o take_suffix Fdl_Lexer.is_digit o raw_explode;
+
+val name_ord = prod_ord string_ord (option_ord int_ord) o
+ pairself (strip_number ##> Int.fromString);
+
+structure VCtab = Table(type key = string val ord = name_ord);
+
+structure VCs = Theory_Data
+(
+ type T =
+ {pfuns: ((string list * string) option * term) Symtab.table,
+ env:
+ {ctxt: Element.context_i list,
+ defs: (binding * thm) list,
+ types: fdl_type tab,
+ funs: (string list * string) tab,
+ ids: (term * string) Symtab.table * Name.context,
+ proving: bool,
+ vcs: (string * bool *
+ (string * expr) list * (string * expr) list) VCtab.table,
+ path: Path.T} option}
+ val empty : T = {pfuns = Symtab.empty, env = NONE}
+ val extend = I
+ fun merge ({pfuns = pfuns1, env = NONE}, {pfuns = pfuns2, env = NONE}) =
+ {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2),
+ env = NONE}
+ | merge _ = err_unfinished ()
+)
+
+fun set_env (env as {funs, ...}) thy = VCs.map (fn
+ {pfuns, env = NONE} =>
+ {pfuns = check_pfuns_types thy funs pfuns, env = SOME env}
+ | _ => err_unfinished ()) thy;
+
+fun mk_pat s = (case Int.fromString s of
+ 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) =
+ let val prop_of =
+ HOLogic.mk_Trueprop o fst o term_of_expr thy types funs pfuns ids
+ in
+ (tr, proved,
+ Element.Assumes (map (fn (s', e) =>
+ ((Binding.name ("H" ^ s'), []), [(prop_of e, [])])) ps),
+ Element.Shows (map (fn (s', e) =>
+ (Attrib.empty_binding, [(prop_of e, mk_pat s')])) cs))
+ end;
+
+fun fold_vcs f vcs =
+ VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs;
+
+fun pfuns_of_vcs pfuns vcs =
+ fold_vcs (add_expr_pfuns o snd) vcs [] |>
+ filter_out (Symtab.defined pfuns);
+
+fun declare_missing_pfuns thy funs pfuns vcs (tab, ctxt) =
+ let
+ val (fs, (tys, Ts)) =
+ pfuns_of_vcs pfuns vcs |>
+ map_filter (fn s => lookup funs s |>
+ Option.map (fn ty => (s, (SOME ty, pfun_type thy ty)))) |>
+ split_list ||> split_list;
+ val (fs', ctxt') = Name.variants fs ctxt
+ in
+ (fold Symtab.update_new (fs ~~ (tys ~~ map Free (fs' ~~ Ts))) pfuns,
+ Element.Fixes (map2 (fn s => fn T =>
+ (Binding.name s, SOME T, NoSyn)) fs' Ts),
+ (tab, ctxt'))
+ end;
+
+fun add_proof_fun prep (s, (optty, raw_t)) thy =
+ VCs.map (fn
+ {env = SOME {proving = true, ...}, ...} => err_unfinished ()
+ | {pfuns, env} =>
+ let
+ val optty' = (case env of
+ SOME {funs, ...} => lookup funs s
+ | NONE => NONE);
+ val optty'' = NONE |> upd_option optty |> upd_option optty';
+ val t = prep (Option.map (pfun_type thy) optty'') raw_t
+ in
+ (check_pfun_type thy s t optty optty';
+ if is_some optty'' orelse is_none env then
+ {pfuns = Symtab.update_new (s, (optty'', t)) pfuns,
+ env = env}
+ handle Symtab.DUP _ => error ("Proof function " ^ s ^
+ " already associated with function")
+ else error ("Undeclared proof function " ^ s))
+ end) thy;
+
+val is_closed = is_none o #env o VCs.get;
+
+fun lookup_vc thy name =
+ (case VCs.get thy of
+ {env = SOME {vcs, types, funs, ids, ctxt, ...}, pfuns} =>
+ (case VCtab.lookup vcs name of
+ 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
+ | NONE => NONE)
+ | _ => NONE);
+
+fun get_vcs thy = (case VCs.get thy of
+ {env = SOME {vcs, types, funs, ids, ctxt, defs, ...}, pfuns} =>
+ let val (pfuns', ctxt', ids') =
+ declare_missing_pfuns thy funs pfuns vcs ids
+ in
+ (ctxt @ [ctxt'], defs,
+ VCtab.dest vcs |>
+ map (apsnd (mk_vc thy types funs pfuns' ids')))
+ end
+ | _ => ([], [], []));
+
+fun mark_proved name = VCs.map (fn
+ {pfuns, env = SOME {ctxt, defs, types, funs, ids, vcs, path, ...}} =>
+ {pfuns = pfuns,
+ env = SOME {ctxt = ctxt, defs = defs,
+ types = types, funs = funs, ids = ids,
+ proving = true,
+ vcs = VCtab.map_entry name (fn (trace, _, ps, cs) =>
+ (trace, true, ps, cs)) vcs,
+ path = path}}
+ | x => x);
+
+fun close thy = VCs.map (fn
+ {pfuns, env = SOME {vcs, path, ...}} =>
+ (case VCtab.fold_rev (fn (s, (_, p, _, _)) =>
+ (if p then apfst else apsnd) (cons s)) vcs ([], []) of
+ (proved, []) =>
+ (File.write (Path.ext "prv" path)
+ (concat (map (fn s => snd (strip_number s) ^
+ " -- proved by " ^ Distribution.version ^ "\n") proved));
+ {pfuns = pfuns, env = NONE})
+ | (_, unproved) => err_vcs unproved)
+ | x => x) thy;
+
+
+(** set up verification conditions **)
+
+fun partition_opt f =
+ let
+ fun part ys zs [] = (rev ys, rev zs)
+ | part ys zs (x :: xs) = (case f x of
+ SOME y => part (y :: ys) zs xs
+ | NONE => part ys (x :: zs) xs)
+ in part [] [] end;
+
+fun dest_def (id, (Substitution_Rule ([], Ident s, rhs))) = SOME (id, (s, rhs))
+ | dest_def _ = NONE;
+
+fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i);
+
+fun add_const (s, ty) ((tab, ctxt), thy) =
+ let
+ val T = mk_type thy ty;
+ val b = Binding.name s;
+ val c = Const (Sign.full_name thy b, T)
+ in
+ (c,
+ ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt),
+ Sign.add_consts_i [(b, T, NoSyn)] thy))
+ end;
+
+fun add_def types funs 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 _ = ty = ty' orelse
+ error ("Declared type " ^ ty ^ " of " ^ s ^
+ "\ndoes not match type " ^ ty' ^ " in definition");
+ val id' = mk_rulename id;
+ val lthy = Named_Target.theory_init thy;
+ val ((t', (_, th)), lthy') = Specification.definition
+ (NONE, ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Free (s, mk_type thy ty), t)))) lthy;
+ val phi = ProofContext.export_morphism lthy' lthy
+ in
+ ((id', Morphism.thm phi th),
+ ((Symtab.update (s, (Morphism.term phi t', ty)) tab,
+ Name.declare s ctxt),
+ Local_Theory.exit_global lthy'))
+ end
+ | NONE => error ("Undeclared constant " ^ s));
+
+fun add_var (s, ty) (ids, thy) =
+ let val ([Free p], ids') = mk_variables thy [s] ty ids
+ in (p, (ids', thy)) end;
+
+fun add_init_vars vcs (ids_thy as ((tab, _), _)) =
+ fold_map add_var
+ (map_filter
+ (fn s => case try (unsuffix "~") s of
+ SOME s' => (case Symtab.lookup tab s' of
+ SOME (_, ty) => SOME (s, ty)
+ | NONE => error ("Undeclared identifier " ^ s'))
+ | NONE => NONE)
+ (fold_vcs (add_expr_idents o snd) vcs []))
+ ids_thy;
+
+fun is_trivial_vc ([], [(_, Ident "true")]) = true
+ | is_trivial_vc _ = false;
+
+fun rulenames rules = commas
+ (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 =
+ (case find_first (fn (_, (_, e)) =>
+ forall (Symtab.defined pfuns) (add_expr_pfuns e []) andalso
+ forall (fn id =>
+ member (fn (s, (_, (s', _))) => s = s') sdefs id orelse
+ member (fn (s, (s', _)) => s = s') consts id)
+ (add_expr_idents e [])) defs of
+ SOME d => sort_defs pfuns consts
+ (remove (op =) d defs) (d :: sdefs)
+ | NONE => error ("Bad definitions: " ^ rulenames defs));
+
+fun set_vcs ({types, vars, consts, funs} : decls) (rules, _) vcs path thy =
+ let
+ val {pfuns, ...} = VCs.get thy;
+ val (defs', rules') = partition_opt dest_def rules;
+ val consts' =
+ subtract (fn ((_, (s, _)), (s', _)) => s = s') defs' (items consts);
+ val defs = sort_defs pfuns consts' defs' [];
+ (* ignore all complex rules in rls files *)
+ val (rules'', other_rules) =
+ List.partition (complex_rule o snd) rules';
+ val _ = if null rules'' then ()
+ else warning ("Ignoring rules: " ^ rulenames rules'');
+
+ val vcs' = VCtab.make (maps (fn (tr, vcs) =>
+ map (fn (s, (ps, cs)) => (s, (tr, false, ps, cs)))
+ (filter_out (is_trivial_vc o snd) vcs)) vcs);
+
+ val _ = (case filter_out (is_some o lookup funs)
+ (pfuns_of_vcs pfuns vcs') of
+ [] => ()
+ | fs => error ("Undeclared proof function(s) " ^ commas fs));
+
+ val (((defs', vars''), ivars), (ids, thy')) =
+ ((Symtab.empty |>
+ Symtab.update ("false", (HOLogic.false_const, booleanN)) |>
+ Symtab.update ("true", (HOLogic.true_const, booleanN)),
+ Name.context), thy) |>
+ fold add_type_def (items types) |>
+ fold (snd oo add_const) consts' |>
+ fold_map (add_def types funs pfuns consts) defs ||>>
+ fold_map add_var (items vars) ||>>
+ add_init_vars 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, [])]))
+ 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'
+ end;
+
+end;