diff --git a/src/HOL/SPARK/Tools/spark_vcs.ML b/src/HOL/SPARK/Tools/spark_vcs.ML --- a/src/HOL/SPARK/Tools/spark_vcs.ML +++ b/src/HOL/SPARK/Tools/spark_vcs.ML @@ -1,1124 +1,1123 @@ (* 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 -> (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 val add_type: string * (typ * (string * string) list) -> theory -> theory val lookup_vc: theory -> bool -> string -> (Element.context_i list * (string * thm list option * Element.context_i * Element.statement_i)) option val get_vcs: theory -> bool -> Element.context_i list * (binding * thm) list * (string * (string * thm list option * Element.context_i * Element.statement_i)) list val mark_proved: string -> thm list -> theory -> theory val close: bool -> theory -> theory val is_closed: theory -> bool end; structure SPARK_VCs: SPARK_VCS = struct open Fdl_Parser; (** theory data **) fun err_unfinished () = error "An unfinished SPARK environment is still open." val strip_number = apply2 implode o chop_suffix Fdl_Lexer.is_digit o raw_explode; val name_ord = prod_ord string_ord (option_ord int_ord) o apply2 (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, type_map: (typ * (string * string) list) Symtab.table, env: {ctxt: Element.context_i list, defs: (binding * thm) list, types: fdl_type tab, funs: (string list * string) tab, pfuns: ((string list * string) option * term) Symtab.table, ids: (term * string) Symtab.table * Name.context, proving: bool, vcs: (string * thm list option * (string * expr) list * (string * expr) list) VCtab.table, 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}, {pfuns = pfuns2, type_map = type_map2, env = NONE}) = {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2), type_map = Symtab.merge (op =) (type_map1, type_map2), env = NONE} | merge _ = err_unfinished () ) (** utilities **) val to_lower = raw_explode #> map Symbol.to_ascii_lower #> implode; val lcase_eq = (op =) o apply2 (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 unprefix_pfun "" s = s | unprefix_pfun prfx s = let val (prfx', s') = strip_prfx s in if prfx = prfx' then s' else s end; 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 get_type thy prfx ty = let val {type_map, ...} = VCs.get thy in lookup_prfx prfx type_map ty end; 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), [])), []) | SOME p => p); fun mk_type thy prfx ty = fst (mk_type' thy prfx ty); val booleanN = "boolean"; val integerN = "integer"; 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 = Proof_Context.init_global (Proof_Context.theory_of lthy'); val thm' = singleton (Proof_Context.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 prfx xs ty (tab, ctxt) = let val T = mk_type thy prfx ty; val (ys, ctxt') = fold_map Name.variant (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; fun get_record_info thy T = (case Record.dest_recTs T of [(tyname, [\<^typ>\unit\])] => Record.get_info thy (Long_Name.qualifier tyname) | _ => NONE); fun find_field [] fname fields = find_first (curry lcase_eq fname o fst) fields | find_field cmap fname fields = (case AList.lookup (op =) cmap fname of NONE => NONE | SOME fname' => SOME (fname', the (AList.lookup (op =) fields fname'))); fun find_field' fname = get_first (fn (flds, fldty) => if member (op =) flds fname then SOME fldty else NONE); fun assoc_ty_err thy T s msg = error ("Type " ^ Syntax.string_of_typ_global thy T ^ " associated with SPARK type " ^ s ^ "\n" ^ msg); (** generate properties of enumeration types **) fun add_enum_type tyname tyname' thy = let val {case_name, ...} = the (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] tyname'); val cs = map Const (the (BNF_LFP_Compat.get_constrs thy tyname')); val k = length cs; val T = Type (tyname', []); 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 |> Class.instantiation ([tyname'], [], \<^sort>\spark_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 {context = ctxt, ...} => resolve_tac ctxt @{thms subset_antisym} 1 THEN resolve_tac ctxt @{thms subsetI} 1 THEN Old_Datatype_Aux.exh_tac ctxt (K (#exhaust (BNF_LFP_Compat.the_info (Proof_Context.theory_of ctxt) [BNF_LFP_Compat.Keep_Nesting] tyname'))) 1 THEN ALLGOALS (asm_full_simp_tac ctxt)); val finite_UNIV = Goal.prove lthy [] [] (HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T)) (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [UNIV_eq]) 1); val card_UNIV = Goal.prove lthy [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (card, HOLogic.mk_number HOLogic.natT k))) (fn {context = ctxt, ...} => simp_tac (ctxt 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 {context = ctxt, ...} => simp_tac (ctxt addsimps [card_UNIV]) 1 THEN simp_tac (ctxt addsimps [UNIV_eq, def1]) 1 THEN resolve_tac ctxt @{thms subset_antisym} 1 THEN simp_tac ctxt 1 THEN resolve_tac ctxt @{thms subsetI} 1 THEN asm_full_simp_tac (ctxt addsimps @{thms interval_expand} delsimps @{thms atLeastLessThan_iff}) 1); val lthy' = Class.prove_instantiation_instance (fn ctxt => Class.intro_classes_tac ctxt [] THEN resolve_tac ctxt [finite_UNIV] 1 THEN resolve_tac ctxt [range_pos] 1 THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def3]) 1 THEN simp_tac (put_simpset HOL_basic_ss ctxt 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 {context = ctxt, ...} => simp_tac (ctxt addsimps [def1]) 1); val th' = Goal.prove lthy' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c))) (fn {context = ctxt, ...} => resolve_tac ctxt [@{thm inj_pos} RS @{thm injD}] 1 THEN simp_tac (ctxt 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 {context = ctxt, ...} => simp_tac (ctxt 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 {context = ctxt, ...} => simp_tac (ctxt addsimps [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1); in lthy' |> Local_Theory.note ((Binding.name (tyname ^ "_card"), @{attributes [simp]}), [card_UNIV]) ||>> Local_Theory.note ((Binding.name (tyname ^ "_pos"), @{attributes [simp]}), pos_eqs) ||>> Local_Theory.note ((Binding.name (tyname ^ "_val"), @{attributes [simp]}), val_eqs) ||>> Local_Theory.note ((Binding.name (tyname ^ "_first_el"), @{attributes [simp]}), [first_el]) ||>> Local_Theory.note ((Binding.name (tyname ^ "_last_el"), @{attributes [simp]}), [last_el]) |> snd |> Local_Theory.exit_global end; 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"); fun check_enum [] [] = NONE | check_enum els [] = SOME ("has no element(s) " ^ commas els) | check_enum [] cs = SOME ("has extra element(s) " ^ commas (map (Long_Name.base_name o fst) cs)) | check_enum (el :: els) ((cname, _) :: cs) = if lcase_eq (el, cname) then check_enum els cs else SOME ("either has no element " ^ el ^ " or it is at the wrong position"); fun invert_map [] = I | invert_map cmap = map (apfst (the o AList.lookup (op =) (map swap cmap))); 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 prfx ty) thy |> snd)) | add_type_def prfx (s, Enum_Type els) ((tab, ctxt), thy) = let val (thy', tyname) = (case get_type thy prfx s of NONE => let val tyb = Binding.name s; val tyname = Sign.full_name thy tyb in (thy |> BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Keep_Nesting] [((tyb, [], NoSyn), map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |> add_enum_type s tyname, tyname) end | SOME (T as Type (tyname, []), cmap) => (case BNF_LFP_Compat.get_constrs thy tyname of NONE => assoc_ty_err thy T s "is not a datatype" | SOME cs => let val (prfx', _) = strip_prfx s in case check_enum (map (unprefix_pfun prfx') els) (invert_map cmap cs) of NONE => (thy, tyname) | SOME msg => assoc_ty_err thy T s msg end) | SOME (T, _) => assoc_ty_err thy T s "is not a datatype"); val cs = map Const (the (BNF_LFP_Compat.get_constrs thy' tyname)); in ((fold (Symtab.update_new o apsnd (rpair s)) (els ~~ cs) tab, fold Name.declare els ctxt), thy') end | 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 prfx) argtys) --> mk_type thy prfx resty) thy |> snd)) | add_type_def prfx (s, Record_Type fldtys) (ids, thy) = (ids, let val fldTs = maps (fn (flds, ty) => map (rpair (mk_type thy prfx ty)) flds) fldtys in case get_type thy prfx s of NONE => Record.add_record {overloaded = false} ([], Binding.name s) NONE (map (fn (fld, T) => (Binding.name fld, T, NoSyn)) fldTs) thy | SOME (rT, cmap) => (case get_record_info thy rT of NONE => assoc_ty_err thy rT s "is not a record type" | SOME {fields, ...} => let val fields' = invert_map cmap fields in (case subtract (lcase_eq o apply2 fst) fldTs fields' of [] => () | flds => assoc_ty_err thy rT s ("has extra field(s) " ^ commas (map (Long_Name.base_name o fst) flds)); map (fn (fld, T) => case AList.lookup lcase_eq fields' fld of NONE => assoc_ty_err thy rT s ("has no field " ^ fld) | SOME U => T = U orelse assoc_ty_err thy rT s ("has field " ^ fld ^ " whose type\n" ^ Syntax.string_of_typ_global thy U ^ "\ndoes not match declared type\n" ^ Syntax.string_of_typ_global thy T)) fldTs; thy) end) end) | add_type_def prfx (s, Pending_Type) (ids, thy) = (ids, case get_type thy prfx s of SOME _ => thy | NONE => Typedecl.typedecl_global {final = true} (Binding.name s, [], NoSyn) thy |> snd); 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) | 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>\modulo\ (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 => (case lookup_prfx prfx pfuns s of SOME (SOME ([], resty), t) => (t, resty) | _ => 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 prfx 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; val (rT, cmap) = 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 cmap fname fields, find_field' fname fldtys) of (SOME (fname', fT), SOME fldty) => (Const (fname', rT --> fT) $ t, fldty) | _ => 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, cmap) = mk_type' thy prfx rcdty; val (u, fldty) = tm_of vs e'; val fT = mk_type thy prfx fldty in case get_record_info thy rT of SOME {fields, ...} => (case find_field cmap fname fields of SOME (fname', fU) => if fT = fU then (Const (fname' ^ "_update", (fT --> fT) --> rT --> rT) $ Abs ("x", fT, u) $ t, rcdty) else error ("Type\n" ^ Syntax.string_of_typ_global thy fT ^ "\ndoes not match type\n" ^ Syntax.string_of_typ_global thy fU ^ "\nof 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 prfx 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 prfx 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 prfx 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 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)))))) | 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 prfx) idxtys); val U = mk_type thy prfx 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)) = let val (T, cmap) = 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") | SOME info => info); val flds' = map (apsnd (tm_of vs)) flds; val fnames = fields |> invert_map cmap |> map (Long_Name.base_name o fst); val fnames' = map fst flds; val (fvals, ftys) = split_list (map (fn s' => case AList.lookup lcase_eq flds' s' of SOME fval_ty => fval_ty | NONE => error ("Field " ^ s' ^ " missing in record " ^ s)) fnames); val _ = (case subtract lcase_eq 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 (ext_name, map (mk_type thy prfx) ftys @ [HOLogic.unitT] ---> T), fvals @ [HOLogic.unit]), s) end | tm_of vs (Array (s, default, assocs)) = (case lookup types s of SOME (Array_Type (idxtys, elty)) => let val Ts = map (mk_type thy prfx) idxtys; val T = foldr1 HOLogic.mk_prodT Ts; 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) $ 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 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)) | 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; fun add_expr_pfuns funs = fold_expr (fn s => if is_pfun s then I else insert (op =) s) (fn s => if is_none (lookup funs s) then I else insert (op =) s); val add_expr_idents = fold_expr (K I) (insert (op =)); fun pfun_type thy prfx (argtys, resty) = map (mk_type thy prfx) argtys ---> mk_type thy prfx resty; 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 prfx 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 prfx funs = Symtab.map (fn s => fn (optty, t) => let val optty' = lookup funs (unprefix_pfun prfx s) in (check_pfun_type thy prfx s t optty optty'; (NONE |> upd_option optty |> upd_option optty', t)) end); (** the VC store **) fun pp_vcs msg vcs = Pretty.big_list msg (map (Pretty.str o fst) vcs); fun pp_open_vcs [] = Pretty.str "All verification conditions have been proved." | pp_open_vcs vcs = pp_vcs "The following verification conditions remain to be proved:" vcs; fun partition_vcs vcs = VCtab.fold_rev (fn (name, (trace, SOME thms, ps, cs)) => apfst (cons (name, (trace, thms, ps, cs))) | (name, (trace, NONE, ps, cs)) => apsnd (cons (name, (trace, ps, cs)))) vcs ([], []); fun insert_break prt = Pretty.blk (0, [Pretty.fbrk, prt]); fun print_open_vcs f vcs = (Pretty.writeln (f (pp_open_vcs (snd (partition_vcs vcs)))); vcs); fun set_env ctxt defs types funs ids vcs path prefix thy = VCs.map (fn {pfuns, type_map, env = NONE} => {pfuns = pfuns, type_map = type_map, env = SOME {ctxt = ctxt, defs = defs, types = types, funs = funs, pfuns = check_pfuns_types thy prefix funs pfuns, ids = ids, proving = false, vcs = print_open_vcs I vcs, path = path, prefix = prefix}} | _ => 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 prfx types pfuns ids name_concl (tr, proved, ps, cs) = let val prop_of = HOLogic.mk_Trueprop o fst o term_of_expr thy prfx types 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) => (if name_concl then (Binding.name ("C" ^ s'), []) else Binding.empty_atts, [(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 prfx funs pfuns vcs = fold_vcs (add_expr_pfuns funs o snd) vcs [] |> filter (is_none o lookup_prfx prfx pfuns); fun declare_missing_pfuns thy prfx funs pfuns vcs (tab, ctxt) = let val (fs, (tys, Ts)) = pfuns_of_vcs prfx funs pfuns vcs |> map_filter (fn s => lookup funs s |> Option.map (fn ty => (s, (SOME ty, pfun_type thy prfx ty)))) |> split_list ||> split_list; val (fs', ctxt') = fold_map Name.variant 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 map_pfuns f {pfuns, type_map, env} = {pfuns = f pfuns, type_map = type_map, env = env} fun map_pfuns_env f {pfuns, type_map, env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env, ids, proving, vcs, path, prefix}} = if proving then err_unfinished () else {pfuns = pfuns, type_map = type_map, env = SOME {ctxt = ctxt, defs = defs, types = types, funs = funs, pfuns = f pfuns_env, ids = ids, proving = false, vcs = vcs, path = path, prefix = prefix}}; fun add_proof_fun prep (s, (optty, raw_t)) thy = VCs.map (fn data as {env, ...} => let val (optty', prfx, map_pf) = (case env of SOME {funs, prefix, ...} => (lookup funs (unprefix_pfun prefix s), prefix, map_pfuns_env) | NONE => (NONE, "", map_pfuns)); val optty'' = NONE |> upd_option optty |> upd_option optty'; 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 [] => () | ts => error ("Term\n" ^ Syntax.string_of_term_global thy t ^ "\nto be associated with proof function " ^ s ^ " contains free variable(s) " ^ commas (map (Syntax.string_of_term_global thy) ts))); in (check_pfun_type thy prfx s t optty optty'; if is_some optty'' orelse is_none env then map_pf (Symtab.update_new (s, (optty'', t))) data handle Symtab.DUP _ => error ("Proof function " ^ s ^ " already associated with function") else error ("Undeclared proof function " ^ s)) end) thy; fun check_mapping _ _ [] _ = () | check_mapping err s cmap cs = (case duplicates (op = o apply2 fst) cmap of [] => (case duplicates (op = o apply2 snd) cmap of [] => (case subtract (op = o apsnd snd) cs cmap of [] => (case subtract (op = o apfst snd) cmap cs of [] => () | zs => err ("has extra " ^ s ^ "(s) " ^ commas zs)) | zs => err ("does not have " ^ s ^ "(s) " ^ commas (map snd zs))) | zs => error ("Several SPARK names are mapped to " ^ commas (map snd zs))) | zs => error ("The SPARK names " ^ commas (map fst zs) ^ " are mapped to more than one name")); fun add_type (s, (T as Type (tyname, Ts), cmap)) thy = let val cmap' = map (apsnd (Sign.intern_const thy)) cmap in thy |> VCs.map (fn {env = SOME _, ...} => err_unfinished () | {pfuns, type_map, env} => {pfuns = pfuns, type_map = Symtab.update_new (s, (T, cmap')) type_map, env = env} handle Symtab.DUP _ => error ("SPARK type " ^ s ^ " already associated with type")) |> (fn thy' => case BNF_LFP_Compat.get_constrs thy' tyname of NONE => (case get_record_info thy' T of NONE => thy' | SOME {fields, ...} => (check_mapping (assoc_ty_err thy' T s) "field" cmap' (map fst fields); thy')) | SOME cs => if null Ts then (map (fn (_, Type (_, [])) => () | (cname, _) => assoc_ty_err thy' T s ("has illegal constructor " ^ Long_Name.base_name cname)) cs; check_mapping (assoc_ty_err thy' T s) "constructor" cmap' (map fst cs); add_enum_type s tyname thy') else assoc_ty_err thy' T s "is illegal") end | add_type (s, (T, _)) thy = assoc_ty_err thy T s "is illegal"; val is_closed = is_none o #env o VCs.get; fun lookup_vc thy name_concl name = (case VCs.get thy of {env = SOME {vcs, types, funs, pfuns, ids, ctxt, prefix, ...}, ...} => (case VCtab.lookup vcs name of SOME vc => let val (pfuns', ctxt', ids') = declare_missing_pfuns thy prefix funs pfuns vcs ids in SOME (ctxt @ [ctxt'], mk_vc thy prefix types pfuns' ids' name_concl vc) end | NONE => NONE) | _ => NONE); fun get_vcs thy name_concl = (case VCs.get thy of {env = SOME {vcs, types, funs, pfuns, ids, ctxt, defs, prefix, ...}, ...} => let val (pfuns', ctxt', ids') = declare_missing_pfuns thy prefix funs pfuns vcs ids in (ctxt @ [ctxt'], defs, VCtab.dest vcs |> map (apsnd (mk_vc thy prefix types pfuns' ids' name_concl))) end | _ => ([], [], [])); fun mark_proved name thms = VCs.map (fn {pfuns, type_map, env = SOME {ctxt, defs, types, funs, pfuns = pfuns_env, ids, vcs, path, prefix, ...}} => {pfuns = pfuns, type_map = type_map, env = SOME {ctxt = ctxt, defs = defs, types = types, funs = funs, pfuns = pfuns_env, ids = ids, proving = true, vcs = print_open_vcs insert_break (VCtab.map_entry name (fn (trace, _, ps, cs) => (trace, SOME thms, ps, cs)) vcs), path = path, prefix = prefix}} | x => x); fun close incomplete thy = thy |> VCs.map (fn {pfuns, type_map, env} => (case env of NONE => error "No SPARK environment is currently open" | SOME {vcs, path, ...} => let val (proved, unproved) = partition_vcs vcs; val _ = Thm.consolidate (maps (#2 o snd) proved); val (proved', proved'') = List.partition (fn (_, (_, thms, _, _)) => Thm_Deps.has_skip_proof thms) proved; val _ = if null unproved then () else (if incomplete then warning else error) (Pretty.string_of (pp_open_vcs unproved)); val _ = if null proved' then () else warning (Pretty.string_of (pp_vcs "The following VCs are not marked as proved \ \because their proofs contain oracles:" proved')); val prv_path = Path.ext "prv" path; val _ = Export.export thy (Path.binding (prv_path, Position.none)) (proved'' |> map (fn (s, _) => XML.Text (snd (strip_number s) ^ - " -- proved by Isabelle/" ^ Isabelle_System.isabelle_id () ^ - Isabelle_System.isabelle_heading () ^ "\n"))); + " -- proved by " ^ Isabelle_System.identification () ^ "\n"))); in {pfuns = pfuns, type_map = type_map, env = NONE} end)) |> Sign.parent_path; (** 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 prfx (s, ty) ((tab, ctxt), thy) = let val T = mk_type thy prfx 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 [(b, T, NoSyn)] thy)) end; 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 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"); val id' = mk_rulename id; val ((t', (_, th)), lthy') = Named_Target.theory_init thy |> Specification.definition NONE [] [] ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq (Free (s, T), t))); val phi = Proof_Context.export_morphism lthy' (Proof_Context.init_global (Proof_Context.theory_of 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 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 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 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 prfx funs pfuns consts defs sdefs = (case find_first (fn (_, (_, e)) => forall (is_some o lookup_prfx prfx pfuns) (add_expr_pfuns funs 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 prfx funs 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 opt_prfx thy = let val prfx' = if opt_prfx = "" then space_implode "__" (Long_Name.explode (Long_Name.qualifier ident)) else opt_prfx; val prfx = to_lower prfx'; val {pfuns, ...} = VCs.get thy; val (defs, rules') = partition_opt dest_def rules; val consts' = subtract (fn ((_, (s, _)), (s', _)) => s = s') defs (items consts); (* 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, NONE, ps, cs))) (filter_out (is_trivial_vc o snd) vcs)) vcs); val _ = (case filter_out (is_some o lookup funs) (pfuns_of_vcs prfx funs pfuns vcs') of [] => () | fs => error ("Undeclared proof function(s) " ^ commas fs)); val (((defs', vars''), ivars), (ids, thy')) = ((Symtab.empty |> Symtab.update ("false", (\<^term>\False\, booleanN)) |> Symtab.update ("true", (\<^term>\True\, booleanN)), Name.context), thy |> Sign.add_path ((if prfx' = "" then "" else prfx' ^ "__") ^ Long_Name.base_name ident)) |> 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 prfx types pfuns consts) (sort_defs prfx funs 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' prfx types pfuns ids rl, [])])) other_rules), Element.Notes ("", [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])] in set_env ctxt defs' types funs ids vcs' path prfx thy' end; end; diff --git a/src/Pure/General/http.scala b/src/Pure/General/http.scala --- a/src/Pure/General/http.scala +++ b/src/Pure/General/http.scala @@ -1,305 +1,304 @@ /* Title: Pure/General/http.scala Author: Makarius HTTP client and server support. */ package isabelle import java.io.{File => JFile} import java.net.{InetSocketAddress, URI, URL, URLConnection, HttpURLConnection} import com.sun.net.httpserver.{HttpExchange, HttpHandler, HttpServer} object HTTP { /** content **/ val mime_type_bytes: String = "application/octet-stream" val mime_type_text: String = "text/plain; charset=utf-8" val mime_type_html: String = "text/html; charset=utf-8" val default_mime_type: String = mime_type_bytes val default_encoding: String = UTF8.charset_name sealed case class Content( bytes: Bytes, file_name: String = "", mime_type: String = default_mime_type, encoding: String = default_encoding, elapsed_time: Time = Time.zero) { def text: String = new String(bytes.array, encoding) } def read_content(file: JFile): Content = { val bytes = Bytes.read(file) val file_name = file.getName val mime_type = Option(URLConnection.guessContentTypeFromName(file_name)).getOrElse(default_mime_type) Content(bytes, file_name = file_name, mime_type = mime_type) } def read_content(path: Path): Content = read_content(path.file) /** client **/ val NEWLINE: String = "\r\n" object Client { val default_timeout: Time = Time.seconds(180) def open_connection(url: URL, timeout: Time = default_timeout, user_agent: String = ""): HttpURLConnection = { url.openConnection match { case connection: HttpURLConnection => if (0 < timeout.ms && timeout.ms <= Integer.MAX_VALUE) { val ms = timeout.ms.toInt connection.setConnectTimeout(ms) connection.setReadTimeout(ms) } proper_string(user_agent).foreach(s => connection.setRequestProperty("User-Agent", s)) connection case _ => error("Bad URL (not HTTP): " + quote(url.toString)) } } def get_content(connection: HttpURLConnection): Content = { val Charset = """.*\bcharset="?([\S^"]+)"?.*""".r val start = Time.now() using(connection.getInputStream)(stream => { val bytes = Bytes.read_stream(stream, hint = connection.getContentLength) val stop = Time.now() val file_name = Url.file_name(connection.getURL) val mime_type = Option(connection.getContentType).getOrElse(default_mime_type) val encoding = (connection.getContentEncoding, mime_type) match { case (enc, _) if enc != null => enc case (_, Charset(enc)) => enc case _ => default_encoding } Content(bytes, file_name = file_name, mime_type = mime_type, encoding = encoding, elapsed_time = stop - start) }) } def get(url: URL, timeout: Time = default_timeout, user_agent: String = ""): Content = get_content(open_connection(url, timeout = timeout, user_agent = user_agent)) def post(url: URL, parameters: List[(String, Any)], timeout: Time = default_timeout, user_agent: String = ""): Content = { val connection = open_connection(url, timeout = timeout, user_agent = user_agent) connection.setRequestMethod("POST") connection.setDoOutput(true) val boundary = UUID.random_string() connection.setRequestProperty( "Content-Type", "multipart/form-data; boundary=" + quote(boundary)) using(connection.getOutputStream)(out => { def output(s: String): Unit = out.write(UTF8.bytes(s)) def output_newline(n: Int = 1): Unit = (1 to n).foreach(_ => output(NEWLINE)) def output_boundary(end: Boolean = false): Unit = output("--" + boundary + (if (end) "--" else "") + NEWLINE) def output_name(name: String): Unit = output("Content-Disposition: form-data; name=" + quote(name)) def output_value(value: Any): Unit = { output_newline(2) output(value.toString) } def output_content(content: Content): Unit = { proper_string(content.file_name).foreach(s => output("; filename=" + quote(s))) output_newline() proper_string(content.mime_type).foreach(s => output("Content-Type: " + s)) output_newline(2) content.bytes.write_stream(out) } output_newline(2) for { (name, value) <- parameters } { output_boundary() output_name(name) value match { case content: Content => output_content(content) case file: JFile => output_content(read_content(file)) case path: Path => output_content(read_content(path)) case _ => output_value(value) } output_newline() } output_boundary(end = true) out.flush() }) get_content(connection) } } /** server **/ /* response */ object Response { def apply( bytes: Bytes = Bytes.empty, content_type: String = mime_type_bytes): Response = new Response(bytes, content_type) val empty: Response = apply() def text(s: String): Response = apply(Bytes(s), mime_type_text) def html(s: String): Response = apply(Bytes(s), mime_type_html) } class Response private[HTTP](val bytes: Bytes, val content_type: String) { override def toString: String = bytes.toString } /* exchange */ class Exchange private[HTTP](val http_exchange: HttpExchange) { def request_method: String = http_exchange.getRequestMethod def request_uri: URI = http_exchange.getRequestURI def read_request(): Bytes = using(http_exchange.getRequestBody)(Bytes.read_stream(_)) def write_response(code: Int, response: Response): Unit = { http_exchange.getResponseHeaders.set("Content-Type", response.content_type) http_exchange.sendResponseHeaders(code, response.bytes.length.toLong) using(http_exchange.getResponseBody)(response.bytes.write_stream) } } /* handler for request method */ sealed case class Arg(method: String, uri: URI, request: Bytes) { def decode_properties: Properties.T = space_explode('&', request.text).map(s => space_explode('=', s) match { case List(a, b) => Url.decode(a) -> Url.decode(b) case _ => error("Malformed key-value pair in HTTP/POST: " + quote(s)) }) } object Handler { def apply(root: String, body: Exchange => Unit): Handler = new Handler(root, (x: HttpExchange) => body(new Exchange(x))) def method(name: String, root: String, body: Arg => Option[Response]): Handler = apply(root, http => { val request = http.read_request() if (http.request_method == name) { val arg = Arg(name, http.request_uri, request) Exn.capture(body(arg)) match { case Exn.Res(Some(response)) => http.write_response(200, response) case Exn.Res(None) => http.write_response(404, Response.empty) case Exn.Exn(ERROR(msg)) => http.write_response(500, Response.text(Output.error_message_text(msg))) case Exn.Exn(exn) => throw exn } } else http.write_response(400, Response.empty) }) def get(root: String, body: Arg => Option[Response]): Handler = method("GET", root, body) def post(root: String, body: Arg => Option[Response]): Handler = method("POST", root, body) } class Handler private(val root: String, val handler: HttpHandler) { override def toString: String = root } /* server */ class Server private[HTTP](val http_server: HttpServer) { def += (handler: Handler): Unit = http_server.createContext(handler.root, handler.handler) def -= (handler: Handler): Unit = http_server.removeContext(handler.root) def start(): Unit = http_server.start() def stop(): Unit = http_server.stop(0) def address: InetSocketAddress = http_server.getAddress def url: String = "http://" + address.getHostName + ":" + address.getPort override def toString: String = url } def server(handlers: List[Handler] = isabelle_resources): Server = { val http_server = HttpServer.create(new InetSocketAddress(isabelle.Server.localhost, 0), 0) http_server.setExecutor(null) val server = new Server(http_server) for (handler <- handlers) server += handler server } /** Isabelle resources **/ lazy val isabelle_resources: List[Handler] = List(welcome(), fonts()) /* welcome */ def welcome(root: String = "/"): Handler = Handler.get(root, arg => if (arg.uri.toString == root) { - val id = Isabelle_System.isabelle_id() - Some(Response.text("Welcome to Isabelle/" + id + Isabelle_System.isabelle_heading())) + Some(Response.text("Welcome to " + Isabelle_System.identification())) } else None) /* fonts */ private lazy val html_fonts: List[Isabelle_Fonts.Entry] = Isabelle_Fonts.fonts(hidden = true) def fonts(root: String = "/fonts"): Handler = Handler.get(root, arg => { val uri_name = arg.uri.toString if (uri_name == root) { Some(Response.text(cat_lines(html_fonts.map(entry => entry.path.file_name)))) } else { html_fonts.collectFirst( { case entry if uri_name == root + "/" + entry.path.file_name => Response(entry.bytes) }) } }) } diff --git a/src/Pure/System/isabelle_system.ML b/src/Pure/System/isabelle_system.ML --- a/src/Pure/System/isabelle_system.ML +++ b/src/Pure/System/isabelle_system.ML @@ -1,136 +1,139 @@ (* Title: Pure/System/isabelle_system.ML Author: Makarius Isabelle system support. *) signature ISABELLE_SYSTEM = sig val bash_process: string -> Process_Result.T val bash_output: string -> string * int val bash: string -> int val bash_functions: unit -> string list val check_bash_function: Proof.context -> string * Position.T -> string val absolute_path: Path.T -> string val make_directory: Path.T -> Path.T val copy_dir: Path.T -> Path.T -> unit val copy_file: Path.T -> Path.T -> unit val copy_file_base: Path.T * Path.T -> Path.T -> unit val create_tmp_path: string -> string -> Path.T val with_tmp_file: string -> string -> (Path.T -> 'a) -> 'a val rm_tree: Path.T -> unit val with_tmp_dir: string -> (Path.T -> 'a) -> 'a val download: string -> Path.T -> unit val isabelle_id: unit -> string val isabelle_identifier: unit -> string option val isabelle_heading: unit -> string val isabelle_name: unit -> string + val identification: unit -> string end; structure Isabelle_System: ISABELLE_SYSTEM = struct (* bash *) fun bash_process script = Scala.function "bash_process" ("export ISABELLE_TMP=" ^ Bash.string (getenv "ISABELLE_TMP") ^ "\n" ^ script) |> split_strings0 |> (fn [] => raise Exn.Interrupt | [err] => error err | a :: b :: c :: d :: lines => let val rc = Value.parse_int a; val (elapsed, cpu) = apply2 (Time.fromMilliseconds o Value.parse_int) (b, c); val (out_lines, err_lines) = chop (Value.parse_int d) lines; in Process_Result.make {rc = rc, out_lines = out_lines, err_lines = err_lines, timing = {elapsed = elapsed, cpu = cpu, gc = Time.zeroTime}} end | _ => raise Fail "Malformed Isabelle/Scala result"); val bash = bash_process #> Process_Result.print #> Process_Result.rc; fun bash_output s = let val res = bash_process s; val _ = warning (Process_Result.err res); in (Process_Result.out res, Process_Result.rc res) end; (* bash functions *) fun bash_functions () = bash_process "declare -Fx" |> Process_Result.check |> Process_Result.out_lines |> map_filter (space_explode " " #> try List.last); fun check_bash_function ctxt arg = Completion.check_entity Markup.bash_functionN (bash_functions () |> map (rpair Position.none)) ctxt arg; (* directory and file operations *) val absolute_path = Path.implode o File.absolute_path; fun scala_function0 name = ignore o Scala.function name o cat_strings0; fun scala_function name = scala_function0 name o map absolute_path; fun make_directory path = (scala_function "make_directory" [path]; path); fun copy_dir src dst = scala_function "copy_dir" [src, dst]; fun copy_file src dst = scala_function "copy_file" [src, dst]; fun copy_file_base (base_dir, src) target_dir = scala_function0 "copy_file_base" [absolute_path base_dir, Path.implode src, absolute_path target_dir]; (* tmp files *) fun create_tmp_path name ext = let val path = File.tmp_path (Path.basic (name ^ serial_string ()) |> Path.ext ext); val _ = File.exists path andalso raise Fail ("Temporary file already exists: " ^ Path.print path); in path end; fun with_tmp_file name ext f = let val path = create_tmp_path name ext in Exn.release (Exn.capture f path before ignore (try File.rm path)) end; (* tmp dirs *) fun rm_tree path = scala_function "rm_tree" [path]; fun with_tmp_dir name f = let val path = create_tmp_path name "" in Exn.release (Exn.capture f (make_directory path) before ignore (try rm_tree path)) end; (* download file *) fun download url file = ignore (Scala.function "download" (cat_strings0 [url, absolute_path file])); (* Isabelle distribution identification *) fun isabelle_id () = Scala.function "isabelle_id" ""; fun isabelle_identifier () = try getenv_strict "ISABELLE_IDENTIFIER"; fun isabelle_heading () = (case isabelle_identifier () of NONE => "" | SOME version => " (" ^ version ^ ")"); fun isabelle_name () = getenv_strict "ISABELLE_NAME"; +fun identification () = "Isabelle/" ^ isabelle_id () ^ isabelle_heading (); + end; diff --git a/src/Pure/System/isabelle_system.scala b/src/Pure/System/isabelle_system.scala --- a/src/Pure/System/isabelle_system.scala +++ b/src/Pure/System/isabelle_system.scala @@ -1,613 +1,615 @@ /* Title: Pure/System/isabelle_system.scala Author: Makarius Fundamental Isabelle system environment: quasi-static module with optional init operation. */ package isabelle import java.io.{File => JFile, IOException} import java.nio.file.{Path => JPath, Files, SimpleFileVisitor, FileVisitResult, StandardCopyOption, FileSystemException} import java.nio.file.attribute.BasicFileAttributes object Isabelle_System { /** bootstrap information **/ def jdk_home(): String = { val java_home = System.getProperty("java.home", "") val home = new JFile(java_home) val parent = home.getParent if (home.getName == "jre" && parent != null && (new JFile(new JFile(parent, "bin"), "javac")).exists) parent else java_home } def bootstrap_directory( preference: String, envar: String, property: String, description: String): String = { val value = proper_string(preference) orElse // explicit argument proper_string(System.getenv(envar)) orElse // e.g. inherited from running isabelle tool proper_string(System.getProperty(property)) getOrElse // e.g. via JVM application boot process error("Unknown " + description + " directory") if ((new JFile(value)).isDirectory) value else error("Bad " + description + " directory " + quote(value)) } /** implicit settings environment **/ abstract class Service @volatile private var _settings: Option[Map[String, String]] = None @volatile private var _services: Option[List[Class[Service]]] = None def settings(): Map[String, String] = { if (_settings.isEmpty) init() // unsynchronized check _settings.get } def services(): List[Class[Service]] = { if (_services.isEmpty) init() // unsynchronized check _services.get } def make_services[C](c: Class[C]): List[C] = for { c1 <- services() if Library.is_subclass(c1, c) } yield c1.getDeclaredConstructor().newInstance().asInstanceOf[C] def init(isabelle_root: String = "", cygwin_root: String = ""): Unit = synchronized { if (_settings.isEmpty || _services.isEmpty) { val isabelle_root1 = bootstrap_directory(isabelle_root, "ISABELLE_ROOT", "isabelle.root", "Isabelle root") val cygwin_root1 = if (Platform.is_windows) bootstrap_directory(cygwin_root, "CYGWIN_ROOT", "cygwin.root", "Cygwin root") else "" if (Platform.is_windows) Cygwin.init(isabelle_root1, cygwin_root1) def set_cygwin_root(): Unit = { if (Platform.is_windows) _settings = Some(_settings.getOrElse(Map.empty) + ("CYGWIN_ROOT" -> cygwin_root1)) } set_cygwin_root() def default(env: Map[String, String], entry: (String, String)): Map[String, String] = if (env.isDefinedAt(entry._1) || entry._2 == "") env else env + entry val env = { val temp_windows = { val temp = if (Platform.is_windows) System.getenv("TEMP") else null if (temp != null && temp.contains('\\')) temp else "" } val user_home = System.getProperty("user.home", "") val isabelle_app = System.getProperty("isabelle.app", "") default( default( default(sys.env + ("ISABELLE_JDK_HOME" -> File.standard_path(jdk_home())), "TEMP_WINDOWS" -> temp_windows), "HOME" -> user_home), "ISABELLE_APP" -> "true") } val settings = { val dump = JFile.createTempFile("settings", null) dump.deleteOnExit try { val cmd1 = if (Platform.is_windows) List(cygwin_root1 + "\\bin\\bash", "-l", File.standard_path(isabelle_root1 + "\\bin\\isabelle")) else List(isabelle_root1 + "/bin/isabelle") val cmd = cmd1 ::: List("getenv", "-d", dump.toString) val (output, rc) = process_output(process(cmd, env = env, redirect = true)) if (rc != 0) error(output) val entries = (for (entry <- Library.split_strings0(File.read(dump)) if entry != "") yield { val i = entry.indexOf('=') if (i <= 0) entry -> "" else entry.substring(0, i) -> entry.substring(i + 1) }).toMap entries + ("PATH" -> entries("PATH_JVM")) - "PATH_JVM" } finally { dump.delete } } _settings = Some(settings) set_cygwin_root() val variable = "ISABELLE_SCALA_SERVICES" val services = for (name <- space_explode(':', settings.getOrElse(variable, getenv_error(variable)))) yield { def err(msg: String): Nothing = error("Bad entry " + quote(name) + " in " + variable + "\n" + msg) try { Class.forName(name).asInstanceOf[Class[Service]] } catch { case _: ClassNotFoundException => err("Class not found") case exn: Throwable => err(Exn.message(exn)) } } _services = Some(services) } } /* getenv -- dynamic process environment */ private def getenv_error(name: String): Nothing = error("Undefined Isabelle environment variable: " + quote(name)) def getenv(name: String, env: Map[String, String] = settings()): String = env.getOrElse(name, "") def getenv_strict(name: String, env: Map[String, String] = settings()): String = proper_string(getenv(name, env)) getOrElse error("Undefined Isabelle environment variable: " + quote(name)) def cygwin_root(): String = getenv_strict("CYGWIN_ROOT") /* getetc -- static distribution parameters */ def getetc(name: String, root: Path = Path.ISABELLE_HOME): Option[String] = { val path = root + Path.basic("etc") + Path.basic(name) if (path.is_file) { Library.trim_split_lines(File.read(path)) match { case Nil => None case List(s) => Some(s) case _ => error("Single line expected in " + path.absolute) } } else None } /* Isabelle distribution identification */ def isabelle_id(root: Path = Path.ISABELLE_HOME): String = getetc("ISABELLE_ID", root = root) orElse Mercurial.archive_id(root) getOrElse { if (Mercurial.is_repository(root)) Mercurial.repository(root).parent() else error("Failed to identify Isabelle distribution " + root) } object Isabelle_Id extends Scala.Fun("isabelle_id") { val here = Scala_Project.here def apply(arg: String): String = isabelle_id() } def isabelle_tags(root: Path = Path.ISABELLE_HOME): String = getetc("ISABELLE_TAGS", root = root) orElse Mercurial.archive_tags(root) getOrElse { if (Mercurial.is_repository(root)) { val hg = Mercurial.repository(root) hg.tags(rev = hg.parent()) } else "" } def isabelle_identifier(): Option[String] = proper_string(getenv("ISABELLE_IDENTIFIER")) def isabelle_heading(): String = isabelle_identifier() match { case None => "" case Some(version) => " (" + version + ")" } def isabelle_name(): String = getenv_strict("ISABELLE_NAME") + def identification(): String = "Isabelle/" + isabelle_id() + isabelle_heading() + /** file-system operations **/ /* scala functions */ private def apply_paths(arg: String, fun: List[Path] => Unit): String = { fun(Library.split_strings0(arg).map(Path.explode)); "" } private def apply_paths1(arg: String, fun: Path => Unit): String = apply_paths(arg, { case List(path) => fun(path) }) private def apply_paths2(arg: String, fun: (Path, Path) => Unit): String = apply_paths(arg, { case List(path1, path2) => fun(path1, path2) }) private def apply_paths3(arg: String, fun: (Path, Path, Path) => Unit): String = apply_paths(arg, { case List(path1, path2, path3) => fun(path1, path2, path3) }) /* permissions */ def chmod(arg: String, path: Path): Unit = bash("chmod " + arg + " " + File.bash_path(path)).check def chown(arg: String, path: Path): Unit = bash("chown " + arg + " " + File.bash_path(path)).check /* directories */ def make_directory(path: Path): Path = { if (!path.is_dir) { try { Files.createDirectories(path.file.toPath) } catch { case ERROR(_) => error("Failed to create directory: " + path.absolute) } } path } def new_directory(path: Path): Path = if (path.is_dir) error("Directory already exists: " + path.absolute) else make_directory(path) def copy_dir(dir1: Path, dir2: Path): Unit = { val res = bash("cp -a " + File.bash_path(dir1) + " " + File.bash_path(dir2)) if (!res.ok) { cat_error("Failed to copy directory " + dir1.absolute + " to " + dir2.absolute, res.err) } } object Make_Directory extends Scala.Fun("make_directory") { val here = Scala_Project.here def apply(arg: String): String = apply_paths1(arg, make_directory) } object Copy_Dir extends Scala.Fun("copy_dir") { val here = Scala_Project.here def apply(arg: String): String = apply_paths2(arg, copy_dir) } /* copy files */ def copy_file(src: JFile, dst: JFile): Unit = { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!File.eq(src, target)) { try { Files.copy(src.toPath, target.toPath, StandardCopyOption.COPY_ATTRIBUTES, StandardCopyOption.REPLACE_EXISTING) } catch { case ERROR(msg) => cat_error("Failed top copy file " + File.path(src).absolute + " to " + File.path(dst).absolute, msg) } } } def copy_file(src: Path, dst: Path): Unit = copy_file(src.file, dst.file) def copy_file_base(base_dir: Path, src: Path, target_dir: Path): Unit = { val src1 = src.expand val src1_dir = src1.dir if (!src1.starts_basic) error("Illegal path specification " + src1 + " beyond base directory") copy_file(base_dir + src1, Isabelle_System.make_directory(target_dir + src1_dir)) } object Copy_File extends Scala.Fun("copy_file") { val here = Scala_Project.here def apply(arg: String): String = apply_paths2(arg, copy_file) } object Copy_File_Base extends Scala.Fun("copy_file_base") { val here = Scala_Project.here def apply(arg: String): String = apply_paths3(arg, copy_file_base) } /* move files */ def move_file(src: JFile, dst: JFile): Unit = { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!File.eq(src, target)) Files.move(src.toPath, target.toPath, StandardCopyOption.REPLACE_EXISTING) } def move_file(src: Path, dst: Path): Unit = move_file(src.file, dst.file) /* symbolic link */ def symlink(src: Path, dst: Path, force: Boolean = false): Unit = { val src_file = src.file val dst_file = dst.file val target = if (dst_file.isDirectory) new JFile(dst_file, src_file.getName) else dst_file if (force) target.delete try { Files.createSymbolicLink(target.toPath, src_file.toPath) } catch { case _: UnsupportedOperationException if Platform.is_windows => Cygwin.link(File.standard_path(src), target) case _: FileSystemException if Platform.is_windows => Cygwin.link(File.standard_path(src), target) } } /* tmp files */ def isabelle_tmp_prefix(): JFile = { val path = Path.explode("$ISABELLE_TMP_PREFIX") path.file.mkdirs // low-level mkdirs to avoid recursion via Isabelle environment File.platform_file(path) } def tmp_file(name: String, ext: String = "", base_dir: JFile = isabelle_tmp_prefix()): JFile = { val suffix = if (ext == "") "" else "." + ext val file = Files.createTempFile(base_dir.toPath, name, suffix).toFile file.deleteOnExit file } def with_tmp_file[A](name: String, ext: String = "")(body: Path => A): A = { val file = tmp_file(name, ext) try { body(File.path(file)) } finally { file.delete } } /* tmp dirs */ def rm_tree(root: JFile): Unit = { root.delete if (root.isDirectory) { Files.walkFileTree(root.toPath, new SimpleFileVisitor[JPath] { override def visitFile(file: JPath, attrs: BasicFileAttributes): FileVisitResult = { try { Files.deleteIfExists(file) } catch { case _: IOException => } FileVisitResult.CONTINUE } override def postVisitDirectory(dir: JPath, e: IOException): FileVisitResult = { if (e == null) { try { Files.deleteIfExists(dir) } catch { case _: IOException => } FileVisitResult.CONTINUE } else throw e } } ) } } def rm_tree(root: Path): Unit = rm_tree(root.file) object Rm_Tree extends Scala.Fun("rm_tree") { val here = Scala_Project.here def apply(arg: String): String = apply_paths1(arg, rm_tree) } def tmp_dir(name: String, base_dir: JFile = isabelle_tmp_prefix()): JFile = { val dir = Files.createTempDirectory(base_dir.toPath, name).toFile dir.deleteOnExit dir } def with_tmp_dir[A](name: String)(body: Path => A): A = { val dir = tmp_dir(name) try { body(File.path(dir)) } finally { rm_tree(dir) } } /* quasi-atomic update of directory */ def update_directory(dir: Path, f: Path => Unit): Unit = { val new_dir = dir.ext("new") val old_dir = dir.ext("old") rm_tree(new_dir) rm_tree(old_dir) f(new_dir) if (dir.is_dir) move_file(dir, old_dir) move_file(new_dir, dir) rm_tree(old_dir) } /** external processes **/ /* raw process */ def process(command_line: List[String], cwd: JFile = null, env: Map[String, String] = settings(), redirect: Boolean = false): Process = { val proc = new ProcessBuilder proc.command(command_line:_*) // fragile on Windows if (cwd != null) proc.directory(cwd) if (env != null) { proc.environment.clear() for ((x, y) <- env) proc.environment.put(x, y) } proc.redirectErrorStream(redirect) proc.start } def process_output(proc: Process): (String, Int) = { proc.getOutputStream.close() val output = File.read_stream(proc.getInputStream) val rc = try { proc.waitFor } finally { proc.getInputStream.close() proc.getErrorStream.close() proc.destroy() Exn.Interrupt.dispose() } (output, rc) } def kill(signal: String, group_pid: String): (String, Int) = { val bash = if (Platform.is_windows) List(cygwin_root() + "\\bin\\bash.exe") else List("/usr/bin/env", "bash") process_output(process(bash ::: List("-c", "kill -" + signal + " -" + group_pid))) } /* GNU bash */ def bash(script: String, cwd: JFile = null, env: Map[String, String] = settings(), redirect: Boolean = false, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), watchdog: Option[Bash.Watchdog] = None, strict: Boolean = true, cleanup: () => Unit = () => ()): Process_Result = { Bash.process(script, cwd = cwd, env = env, redirect = redirect, cleanup = cleanup). result(progress_stdout = progress_stdout, progress_stderr = progress_stderr, watchdog = watchdog, strict = strict) } private lazy val gnutar_check: Boolean = try { bash("tar --version").check.out.containsSlice("GNU tar") || error("") } catch { case ERROR(_) => false } def gnutar( args: String, dir: Path = Path.current, original_owner: Boolean = false, redirect: Boolean = false): Process_Result = { val options = (if (dir.is_current) "" else "-C " + File.bash_path(dir) + " ") + (if (original_owner) "" else "--owner=root --group=staff ") if (gnutar_check) bash("tar " + options + args, redirect = redirect) else error("Expected to find GNU tar executable") } def require_command(cmds: String*): Unit = { for (cmd <- cmds) { if (!bash(Bash.string(cmd) + " --version").ok) error("Missing system command: " + quote(cmd)) } } def hostname(): String = bash("hostname -s").check.out def open(arg: String): Unit = bash("exec \"$ISABELLE_OPEN\" " + Bash.string(arg) + " >/dev/null 2>/dev/null &") def pdf_viewer(arg: Path): Unit = bash("exec \"$PDF_VIEWER\" " + File.bash_path(arg) + " >/dev/null 2>/dev/null &") def open_external_file(name: String): Boolean = { val ext = Library.take_suffix((c: Char) => c != '.', name.toList)._2.mkString val external = ext.nonEmpty && Library.space_explode(':', getenv("ISABELLE_EXTERNAL_FILES")).contains(ext) if (external) { if (ext == "pdf" && Path.is_wellformed(name)) pdf_viewer(Path.explode(name)) else open(name) } external } def export_isabelle_identifier(isabelle_identifier: String): String = if (isabelle_identifier == "") "" else "export ISABELLE_IDENTIFIER=" + Bash.string(isabelle_identifier) + "\n" /** Isabelle resources **/ /* repository clone with Admin */ def admin(): Boolean = Path.explode("~~/Admin").is_dir /* components */ def components(): List[Path] = Path.split(getenv_strict("ISABELLE_COMPONENTS")) /* default logic */ def default_logic(args: String*): String = { args.find(_ != "") match { case Some(logic) => logic case None => getenv_strict("ISABELLE_LOGIC") } } /* download file */ def download(url_name: String, file: Path, progress: Progress = new Progress): Unit = { val url = Url(url_name) progress.echo("Getting " + quote(url_name)) val content = try { HTTP.Client.get(url) } catch { case ERROR(msg) => cat_error("Failed to download " + quote(url_name), msg) } Bytes.write(file, content.bytes) } object Download extends Scala.Fun("download", thread = true) { val here = Scala_Project.here def apply(arg: String): String = Library.split_strings0(arg) match { case List(url, file) => download(url, Path.explode(file)); "" } } }