diff --git a/lib/Tools/version b/lib/Tools/version --- a/lib/Tools/version +++ b/lib/Tools/version @@ -1,113 +1,113 @@ #!/usr/bin/env bash # # Author: Stefan Berghofer, TU Muenchen # Author: Makarius # # DESCRIPTION: display Isabelle version PRG="$(basename "$0")" function usage() { echo echo "Usage: isabelle $PRG [OPTIONS]" echo echo " Options are:" echo " -i short identification (derived from Mercurial id)" echo " -t symbolic tags (derived from Mercurial id)" echo echo " Display Isabelle version information." echo exit 1 } function fail() { echo "$1" >&2 exit 2 } ## process command line # options SHORT_ID="" TAGS="" while getopts "it" OPT do case "$OPT" in i) SHORT_ID=true ;; t) TAGS=true ;; \?) usage ;; esac done shift $(($OPTIND - 1)) # args [ "$#" -ne 0 ] && usage ## main if [ -z "$SHORT_ID" -a -z "$TAGS" ]; then - echo 'repository version' # filled in automatically! + echo "$ISABELLE_NAME" fi HG_ARCHIVAL="$ISABELLE_HOME/.hg_archival.txt" export LANG=C export HGPLAIN= if [ -n "$SHORT_ID" ]; then if [ -f "$ISABELLE_HOME/etc/ISABELLE_ID" ]; then RESULT="$(cat "$ISABELLE_HOME/etc/ISABELLE_ID")" RC="$?" elif [ -d "$ISABELLE_HOME/.hg" ]; then RESULT=$("${HG:-hg}" -R "$ISABELLE_HOME" log -r "p1()" --template="{node|short}\n" 2>/dev/null) RC="$?" elif [ -f "$HG_ARCHIVAL" ]; then RESULT="$(grep "^node:" < "$HG_ARCHIVAL" | cut -d " " -f2 | head -c12)" RC="$?" else RESULT="" RC="0" fi if [ "$RC" -ne 0 ]; then exit "$RC" elif [ -n "$RESULT" ]; then echo "$RESULT" fi fi if [ -n "$TAGS" ]; then if [ -f "$ISABELLE_HOME/etc/ISABELLE_TAGS" ]; then RESULT="$(cat "$ISABELLE_HOME/etc/ISABELLE_TAGS")" RC="$?" elif [ -d "$ISABELLE_HOME/.hg" ]; then RESULT=$("${HG:-hg}" -R "$ISABELLE_HOME" id -t 2>/dev/null) RC="$?" elif [ -f "$HG_ARCHIVAL" ]; then RESULT="$(grep "^tag:" < "$HG_ARCHIVAL" | cut -d " " -f2)" RC="$?" else RESULT="" RC="0" fi if [ "$RC" -ne 0 ]; then exit "$RC" elif [ -n "$RESULT" -a "$RESULT" != "tip" ]; then echo "$RESULT" fi fi diff --git a/lib/scripts/getsettings b/lib/scripts/getsettings --- a/lib/scripts/getsettings +++ b/lib/scripts/getsettings @@ -1,130 +1,134 @@ # -*- shell-script -*- :mode=shellscript: # # Author: Makarius # # Static Isabelle environment for root of process tree. export ISABELLE_HOME export BASH_ENV="$ISABELLE_HOME/lib/scripts/getfunctions" source "$BASH_ENV" if [ -z "$ISABELLE_SETTINGS_PRESENT" ] then export ISABELLE_SETTINGS_PRESENT=true set -o allexport #sane environment defaults (notably on macOS) if [ "$ISABELLE_APP" = true -a -x /usr/libexec/path_helper ]; then eval $(/usr/libexec/path_helper -s) fi #Cygwin vs. POSIX if [ "$OSTYPE" = cygwin ] then unset INI_DIR if [ -n "$TEMP_WINDOWS" ]; then TMPDIR="$(cygpath -u "$TEMP_WINDOWS")" TMP="$TMPDIR" TEMP="$TMPDIR" fi if [ -z "$USER_HOME" ]; then USER_HOME="$(cygpath -u "$USERPROFILE")" fi CYGWIN_ROOT="$(platform_path "/")" ISABELLE_ROOT="$(platform_path "$ISABELLE_HOME")" ISABELLE_CLASSPATH="$(cygpath -i -u -p "$CLASSPATH")" unset CLASSPATH else if [ -z "$USER_HOME" ]; then USER_HOME="$HOME" fi ISABELLE_ROOT="$ISABELLE_HOME" ISABELLE_CLASSPATH="$CLASSPATH" unset CLASSPATH fi #init cumulative settings ISABELLE_FONTS="" ISABELLE_FONTS_HIDDEN="" ISABELLE_SCALA_SERVICES="" ISABELLE_DIRECTORIES="" #main executables ISABELLE_TOOL="$ISABELLE_HOME/bin/isabelle" ISABELLE_SCALA_SCRIPT="$ISABELLE_HOME/bin/isabelle_scala_script" PATH="$ISABELLE_HOME/bin:$PATH" #platform source "$ISABELLE_HOME/lib/scripts/isabelle-platform" if [ -z "$ISABELLE_PLATFORM_FAMILY" ]; then echo 1>&2 "Failed to determine hardware and operating system type!" exit 2 fi -#Isabelle distribution identifier -- filled in automatically! -[ -z "$ISABELLE_IDENTIFIER" ] && ISABELLE_IDENTIFIER="" +if [ -z "$ISABELLE_IDENTIFIER" -a -f "$ISABELLE_HOME/etc/ISABELLE_IDENTIFIER" ] +then + ISABELLE_IDENTIFIER="$(cat "$ISABELLE_HOME/etc/ISABELLE_IDENTIFIER")" +fi + +ISABELLE_NAME="${ISABELLE_IDENTIFIER:-Isabelle}" # components ISABELLE_COMPONENTS="" ISABELLE_COMPONENTS_MISSING="" #main components init_component "$ISABELLE_HOME" [ -d "$ISABELLE_HOME/Admin" ] && init_component "$ISABELLE_HOME/Admin" if [ -d "$ISABELLE_HOME_USER" ]; then init_component "$ISABELLE_HOME_USER" else mkdir -p "$ISABELLE_HOME_USER" chmod $(umask -S) "$ISABELLE_HOME_USER" fi #POLYML_EXE if [ "$ISABELLE_PLATFORM_FAMILY" = "windows" ]; then POLYML_EXE="$ML_HOME/poly.exe" else POLYML_EXE="$ML_HOME/poly" fi #ML system identifier if [ -z "$ML_PLATFORM" ]; then ML_IDENTIFIER="$ML_SYSTEM" else ML_IDENTIFIER="${ML_SYSTEM}_${ML_PLATFORM}" fi #enforce ISABELLE_OCAMLFIND if [ -d "$ISABELLE_OPAM_ROOT/$ISABELLE_OCAML_VERSION/bin" ]; then ISABELLE_OCAMLFIND="$ISABELLE_HOME/lib/scripts/ocamlfind" fi #enforce ISABELLE_GHC if [ -f "$ISABELLE_STACK_ROOT/ISABELLE_GHC_EXE-$ISABELLE_PLATFORM_FAMILY" ]; then if [ -f "$(cat "$ISABELLE_STACK_ROOT/ISABELLE_GHC_EXE-$ISABELLE_PLATFORM_FAMILY")" ]; then ISABELLE_GHC="$ISABELLE_HOME/lib/scripts/ghc" ISABELLE_GHC_STACK=true fi fi #enforce JAVA_HOME if [ -d "$ISABELLE_JDK_HOME/jre" ] then export JAVA_HOME="$ISABELLE_JDK_HOME/jre" else export JAVA_HOME="$ISABELLE_JDK_HOME" fi set +o allexport fi 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,1122 +1,1124 @@ (* 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 " ^ Distribution.version ^ "\n"))); + XML.Text (snd (strip_number s) ^ + " -- proved by Isabelle/" ^ Isabelle_System.isabelle_id () ^ + Isabelle_System.isabelle_heading () ^ "\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/Admin/build_release.scala b/src/Pure/Admin/build_release.scala --- a/src/Pure/Admin/build_release.scala +++ b/src/Pure/Admin/build_release.scala @@ -1,927 +1,908 @@ /* Title: Pure/Admin/build_release.scala Author: Makarius Build full Isabelle distribution from repository. */ package isabelle object Build_Release { /** release info **/ sealed case class Bundle_Info( platform: Platform.Family.Value, platform_description: String, name: String) { def path: Path = Path.explode(name) } class Release private[Build_Release]( progress: Progress, val date: Date, val dist_name: String, val dist_dir: Path, val dist_version: String, val ident: String, val tags: String) { val isabelle: Path = Path.explode(dist_name) val isabelle_dir: Path = dist_dir + isabelle val isabelle_id: Path = isabelle_dir + Path.explode("etc/ISABELLE_ID") val isabelle_tags: Path = isabelle_dir + Path.explode("etc/ISABELLE_TAGS") + val isabelle_identifier: Path = isabelle_dir + Path.explode("etc/ISABELLE_IDENTIFIER") val isabelle_archive: Path = dist_dir + Path.explode(dist_name + ".tar.gz") val isabelle_library_archive: Path = dist_dir + Path.explode(dist_name + "_library.tar.gz") def other_isabelle(dir: Path): Other_Isabelle = Other_Isabelle(dir + isabelle, isabelle_identifier = dist_name + "-build", progress = progress) def bundle_info(platform: Platform.Family.Value): Bundle_Info = platform match { case Platform.Family.linux => Bundle_Info(platform, "Linux", dist_name + "_linux.tar.gz") case Platform.Family.macos => Bundle_Info(platform, "macOS", dist_name + "_macos.tar.gz") case Platform.Family.windows => Bundle_Info(platform, "Windows", dist_name + ".exe") } } /** generated content **/ /* patch release */ - private val getsettings_path = Path.explode("lib/scripts/getsettings") - def patch_release(release: Release): Unit = { val dir = release.isabelle_dir - for (name <- List("src/Pure/System/distribution.ML", "src/Pure/System/distribution.scala")) - { - File.change(dir + Path.explode(name), - _.replace("val is_identified = false", "val is_identified = true")) - } - - File.change(dir + getsettings_path, - _.replace("ISABELLE_IDENTIFIER=\"\"", "ISABELLE_IDENTIFIER=" + quote(release.dist_name))) - File.change(dir + Path.explode("lib/html/library_index_header.template"), _.replace("{ISABELLE}", release.dist_name)) - for { - name <- - List( - "src/Pure/System/distribution.ML", - "src/Pure/System/distribution.scala", - "lib/Tools/version") } - { - File.change(dir + Path.explode(name), _.replace("repository version", release.dist_version)) - } - File.change(dir + Path.explode("README"), _.replace("some repository version of Isabelle", release.dist_version)) } /* ANNOUNCE */ def make_announce(release: Release): Unit = { File.write(release.isabelle_dir + Path.explode("ANNOUNCE"), """ IMPORTANT NOTE ============== This is a snapshot of Isabelle/""" + release.ident + """ from the repository. """) } /* NEWS */ def make_news(other_isabelle: Other_Isabelle, dist_version: String): Unit = { val target = other_isabelle.isabelle_home + Path.explode("doc") val target_fonts = Isabelle_System.make_directory(target + Path.explode("fonts")) other_isabelle.copy_fonts(target_fonts) HTML.write_document(target, "NEWS.html", List(HTML.title("NEWS (" + dist_version + ")")), List( HTML.chapter("NEWS"), HTML.source( Symbol.decode(File.read(other_isabelle.isabelle_home + Path.explode("NEWS")))))) } /* bundled components */ class Bundled(platform: Option[Platform.Family.Value] = None) { def detect(s: String): Boolean = s.startsWith("#bundled") && !s.startsWith("#bundled ") def apply(name: String): String = "#bundled" + (platform match { case None => "" case Some(plat) => "-" + plat }) + ":" + name private val Pattern1 = ("""^#bundled:(.*)$""").r private val Pattern2 = ("""^#bundled-(.*):(.*)$""").r def unapply(s: String): Option[String] = s match { case Pattern1(name) => Some(name) case Pattern2(Platform.Family(plat), name) if platform == Some(plat) => Some(name) case _ => None } } def record_bundled_components(dir: Path): Unit = { val catalogs = List("main", "bundled").map((_, new Bundled())) ::: default_platform_families.flatMap(platform => List(platform.toString, "bundled-" + platform.toString). map((_, new Bundled(platform = Some(platform))))) File.append(Components.components(dir), terminate_lines("#bundled components" :: (for { (catalog, bundled) <- catalogs.iterator path = Components.admin(dir) + Path.basic(catalog) if path.is_file line <- split_lines(File.read(path)) if line.nonEmpty && !line.startsWith("#") && !line.startsWith("jedit_build") } yield bundled(line)).toList)) } def get_bundled_components(dir: Path, platform: Platform.Family.Value): (List[String], String) = { val Bundled = new Bundled(platform = Some(platform)) val components = for { Bundled(name) <- Components.read_components(dir) if !name.startsWith("jedit_build") } yield name val jdk_component = components.find(_.startsWith("jdk")) getOrElse error("Missing jdk component") (components, jdk_component) } def activate_components( dir: Path, platform: Platform.Family.Value, more_names: List[String]): Unit = { def contrib_name(name: String): String = Components.contrib(name = name).implode val Bundled = new Bundled(platform = Some(platform)) Components.write_components(dir, Components.read_components(dir).flatMap(line => line match { case Bundled(name) => if (Components.check_dir(Components.contrib(dir, name))) Some(contrib_name(name)) else None case _ => if (Bundled.detect(line)) None else Some(line) }) ::: more_names.map(contrib_name)) } def make_contrib(dir: Path): Unit = { Isabelle_System.make_directory(Components.contrib(dir)) File.write(Components.contrib(dir, "README"), """This directory contains add-on components that contribute to the main Isabelle distribution. Separate licensing conditions apply, see each directory individually. """) } /** build release **/ private def execute(dir: Path, script: String): Unit = Isabelle_System.bash(script, cwd = dir.file).check private def execute_tar(dir: Path, args: String): Unit = Isabelle_System.gnutar(args, dir = dir).check /* build heaps on remote server */ private def remote_build_heaps( options: Options, platform: Platform.Family.Value, build_sessions: List[String], local_dir: Path): Unit = { val server_option = "build_host_" + platform.toString options.string(server_option) match { case SSH.Target(user, host) => using(SSH.open_session(options, host = host, user = user))(ssh => { Isabelle_System.with_tmp_file("tmp", "tar")(local_tmp_tar => { execute_tar(local_dir, "-cf " + File.bash_path(local_tmp_tar) + " .") ssh.with_tmp_dir(remote_dir => { val remote_tmp_tar = remote_dir + Path.basic("tmp.tar") ssh.write_file(remote_tmp_tar, local_tmp_tar) val remote_commands = List( "cd " + File.bash_path(remote_dir), "tar -xf tmp.tar", "./bin/isabelle build -o system_heaps -b -- " + Bash.strings(build_sessions), "tar -cf tmp.tar heaps") ssh.execute(remote_commands.mkString(" && ")).check ssh.read_file(remote_tmp_tar, local_tmp_tar) }) execute_tar(local_dir, "-xf " + File.bash_path(local_tmp_tar)) }) }) case s => error("Bad " + server_option + ": " + quote(s)) } } /* Isabelle application */ def make_isabelle_options(path: Path, options: List[String], line_ending: String = "\n"): Unit = { val title = "# Java runtime options" File.write(path, (title :: options).map(_ + line_ending).mkString) } def make_isabelle_app( platform: Platform.Family.Value, isabelle_target: Path, isabelle_name: String, jdk_component: String, classpath: List[Path], dock_icon: Boolean = false): Unit = { val script = """#!/usr/bin/env bash # # Author: Makarius # # Main Isabelle application script. # minimal Isabelle environment ISABELLE_HOME="$(cd "$(dirname "$0")"; cd "$(pwd -P)/../.."; pwd)" source "$ISABELLE_HOME/lib/scripts/isabelle-platform" #paranoia settings -- avoid intrusion of alien options unset "_JAVA_OPTIONS" unset "JAVA_TOOL_OPTIONS" #paranoia settings -- avoid problems of Java/Swing versus XIM/IBus etc. unset XMODIFIERS COMPONENT="$ISABELLE_HOME/contrib/""" + jdk_component + """" source "$COMPONENT/etc/settings" # main declare -a JAVA_OPTIONS=($(perl -p -e 's,#.*$,,g;' "$ISABELLE_HOME/Isabelle.options")) "$ISABELLE_HOME/bin/isabelle" env "$ISABELLE_HOME/lib/scripts/java-gui-setup" exec "$ISABELLE_JDK_HOME/bin/java" \ "-Disabelle.root=$ISABELLE_HOME" "${JAVA_OPTIONS[@]}" \ -classpath """" + classpath.map(p => "$ISABELLE_HOME/" + p.implode).mkString(":") + """" \ "-splash:$ISABELLE_HOME/lib/logo/isabelle.gif" \ """ + (if (dock_icon) """"-Xdock:icon=$ISABELLE_HOME/lib/logo/isabelle_transparent-128.png" \ """ else "") + """isabelle.Main "$@" """ val script_path = isabelle_target + Path.explode("lib/scripts/Isabelle_app") File.write(script_path, script) File.set_executable(script_path, true) val component_dir = isabelle_target + Path.explode("contrib/Isabelle_app") Isabelle_System.move_file( component_dir + Path.explode(Platform.standard_platform(platform)) + Path.explode("Isabelle"), isabelle_target + Path.explode(isabelle_name)) Isabelle_System.rm_tree(component_dir) } def make_isabelle_plist(path: Path, isabelle_name: String, isabelle_rev: String): Unit = { File.write(path, """ CFBundleDevelopmentRegion English CFBundleIconFile isabelle.icns CFBundleIdentifier de.tum.in.isabelle CFBundleDisplayName """ + isabelle_name + """ CFBundleInfoDictionaryVersion 6.0 CFBundleName """ + isabelle_name + """ CFBundlePackageType APPL CFBundleShortVersionString """ + isabelle_name + """ CFBundleSignature ???? CFBundleVersion """ + isabelle_rev + """ NSHumanReadableCopyright LSMinimumSystemVersion 10.11 LSApplicationCategoryType public.app-category.developer-tools NSHighResolutionCapable true NSSupportsAutomaticGraphicsSwitching true CFBundleDocumentTypes CFBundleTypeExtensions thy CFBundleTypeIconFile theory.icns CFBundleTypeName Isabelle theory file CFBundleTypeRole Editor LSTypeIsPackage """) } /* main */ private val default_platform_families: List[Platform.Family.Value] = List(Platform.Family.linux, Platform.Family.windows, Platform.Family.macos) def build_release(base_dir: Path, options: Options, components_base: Path = Components.default_components_base, progress: Progress = new Progress, rev: String = "", afp_rev: String = "", proper_release_name: Option[String] = None, platform_families: List[Platform.Family.Value] = default_platform_families, more_components: List[Path] = Nil, website: Option[Path] = None, build_sessions: List[String] = Nil, build_library: Boolean = false, parallel_jobs: Int = 1): Release = { val hg = Mercurial.repository(Path.ISABELLE_HOME) val release = { val date = Date.now() val dist_name = proper_release_name getOrElse ("Isabelle_" + Date.Format.date(date)) val dist_dir = (base_dir + Path.explode("dist-" + dist_name)).absolute val version = proper_string(rev) orElse proper_release_name getOrElse "tip" val ident = try { hg.id(version) } catch { case ERROR(msg) => cat_error("Bad repository version: " + version, msg) } val tags = hg.tags(rev = ident) val dist_version = proper_release_name match { case Some(name) => name + ": " + Date.Format("LLLL uuuu")(date) case None => "Isabelle repository snapshot " + ident + " " + Date.Format.date(date) } new Release(progress, date, dist_name, dist_dir, dist_version, ident, tags) } /* make distribution */ if (release.isabelle_archive.is_file) { progress.echo_warning("Release archive already exists: " + release.isabelle_archive) val archive_ident = Isabelle_System.with_tmp_dir("build_release")(tmp_dir => { - val getsettings = release.isabelle + getsettings_path + val getsettings = release.isabelle + Path.explode("lib/scripts/getsettings") execute_tar(tmp_dir, "-xzf " + File.bash_path(release.isabelle_archive) + " " + File.bash_path(getsettings)) Isabelle_System.isabelle_id(root = tmp_dir + release.isabelle) .getOrElse(error("Failed to determine ISABELLE_ID from " + release.isabelle_archive)) }) if (release.ident != archive_ident) { error("Mismatch of release identification " + release.ident + " vs. archive " + archive_ident) } } else { progress.echo_warning("Producing release archive " + release.isabelle_archive + " ...") Isabelle_System.make_directory(release.dist_dir) if (release.isabelle_dir.is_dir) error("Directory " + release.isabelle_dir + " already exists") progress.echo_warning("Retrieving Mercurial repository version " + release.ident) hg.archive(release.isabelle_dir.expand.implode, rev = release.ident, options = "--type files") for (name <- List(".hg_archival.txt", ".hgtags", ".hgignore", "README_REPOSITORY")) { (release.isabelle_dir + Path.explode(name)).file.delete } progress.echo_warning("Preparing distribution " + quote(release.dist_name)) File.write(release.isabelle_id, release.ident) File.write(release.isabelle_tags, release.tags) + File.write(release.isabelle_identifier, release.dist_name) patch_release(release) if (proper_release_name.isEmpty) make_announce(release) make_contrib(release.isabelle_dir) execute(release.isabelle_dir, """find . -print | xargs chmod -f u+rw""") record_bundled_components(release.isabelle_dir) /* build tools and documentation */ val other_isabelle = release.other_isabelle(release.dist_dir) other_isabelle.init_settings( other_isabelle.init_components(components_base = components_base, catalogs = List("main"))) other_isabelle.resolve_components(echo = true) try { val export_classpath = "export CLASSPATH=" + Bash.string(other_isabelle.getenv("ISABELLE_CLASSPATH")) + "\n" other_isabelle.bash(export_classpath + "./Admin/build all", echo = true).check other_isabelle.bash(export_classpath + "./bin/isabelle jedit -b", echo = true).check } catch { case ERROR(msg) => cat_error("Failed to build tools:", msg) } try { other_isabelle.bash( "./bin/isabelle build_doc -a -o system_heaps -j " + parallel_jobs, echo = true).check } catch { case ERROR(msg) => cat_error("Failed to build documentation:", msg) } make_news(other_isabelle, release.dist_version) for (name <- List("Admin", "browser_info", "heaps")) { Isabelle_System.rm_tree(other_isabelle.isabelle_home + Path.explode(name)) } other_isabelle.cleanup() progress.echo_warning("Creating distribution archive " + release.isabelle_archive) def execute_dist_name(script: String): Unit = Isabelle_System.bash(script, cwd = release.dist_dir.file, env = Isabelle_System.settings() + ("DIST_NAME" -> release.dist_name)).check execute_dist_name(""" set -e chmod -R a+r "$DIST_NAME" chmod -R u+w "$DIST_NAME" chmod -R g=o "$DIST_NAME" find "$DIST_NAME" -type f "(" -name "*.thy" -o -name "*.ML" -o -name "*.scala" ")" -print | xargs chmod -f u-w """) execute_tar(release.dist_dir, "-czf " + File.bash_path(release.isabelle_archive) + " " + Bash.string(release.dist_name)) execute_dist_name(""" set -e mv "$DIST_NAME" "${DIST_NAME}-old" mkdir "$DIST_NAME" mv "${DIST_NAME}-old/README" "${DIST_NAME}-old/NEWS" "${DIST_NAME}-old/ANNOUNCE" \ "${DIST_NAME}-old/COPYRIGHT" "${DIST_NAME}-old/CONTRIBUTORS" "$DIST_NAME" mkdir "$DIST_NAME/doc" mv "${DIST_NAME}-old/doc/"*.pdf \ "${DIST_NAME}-old/doc/"*.html \ "${DIST_NAME}-old/doc/"*.css \ "${DIST_NAME}-old/doc/fonts" \ "${DIST_NAME}-old/doc/Contents" "$DIST_NAME/doc" rm -f Isabelle && ln -sf "$DIST_NAME" Isabelle rm -rf "${DIST_NAME}-old" """) } /* make application bundles */ val bundle_infos = platform_families.map(release.bundle_info) for (bundle_info <- bundle_infos) { val isabelle_name = release.dist_name val platform = bundle_info.platform progress.echo("\nApplication bundle for " + platform) Isabelle_System.with_tmp_dir("build_release")(tmp_dir => { // release archive execute_tar(tmp_dir, "-xzf " + File.bash_path(release.isabelle_archive)) val other_isabelle = release.other_isabelle(tmp_dir) val isabelle_target = other_isabelle.isabelle_home // bundled components progress.echo("Bundled components:") val contrib_dir = Components.contrib(isabelle_target) val (bundled_components, jdk_component) = get_bundled_components(isabelle_target, platform) Components.resolve(components_base, bundled_components, target_dir = Some(contrib_dir), copy_dir = Some(release.dist_dir + Path.explode("contrib")), progress = progress) val more_components_names = more_components.map(Components.unpack(contrib_dir, _, progress = progress)) Components.purge(contrib_dir, platform) activate_components(isabelle_target, platform, more_components_names) // Java parameters val java_options: List[String] = (for { variable <- List( "ISABELLE_JAVA_SYSTEM_OPTIONS", "JEDIT_JAVA_SYSTEM_OPTIONS", "JEDIT_JAVA_OPTIONS") opt <- Word.explode(other_isabelle.getenv(variable)) } yield { val s = "-Dapple.awt.application.name=" if (opt.startsWith(s)) s + isabelle_name else opt }) ::: List("-Disabelle.jedit_server=" + isabelle_name) val classpath: List[Path] = { val base = isabelle_target.absolute Path.split(other_isabelle.getenv("ISABELLE_CLASSPATH")).map(path => { val abs_path = path.absolute File.relative_path(base, abs_path) match { case Some(rel_path) => rel_path case None => error("Bad ISABELLE_CLASSPATH element: " + abs_path) } }) ::: List(Path.explode("src/Tools/jEdit/dist/jedit.jar")) } val jedit_options = Path.explode("src/Tools/jEdit/etc/options") val jedit_props = Path.explode("src/Tools/jEdit/dist/properties/jEdit.props") // build heaps if (build_sessions.nonEmpty) { progress.echo("Building heaps ...") remote_build_heaps(options, platform, build_sessions, isabelle_target) } // application bundling platform match { case Platform.Family.linux => File.change(isabelle_target + jedit_options, _.replaceAll("jedit_reset_font_size : int =.*", "jedit_reset_font_size : int = 24")) File.change(isabelle_target + jedit_props, _.replaceAll("console.fontsize=.*", "console.fontsize=18") .replaceAll("helpviewer.fontsize=.*", "helpviewer.fontsize=18") .replaceAll("metal.primary.fontsize=.*", "metal.primary.fontsize=18") .replaceAll("metal.secondary.fontsize=.*", "metal.secondary.fontsize=18") .replaceAll("view.fontsize=.*", "view.fontsize=24") .replaceAll("view.gutter.fontsize=.*", "view.gutter.fontsize=16")) make_isabelle_options( isabelle_target + Path.explode("Isabelle.options"), java_options) make_isabelle_app(platform, isabelle_target, isabelle_name, jdk_component, classpath) val archive_name = isabelle_name + "_linux.tar.gz" progress.echo("Packaging " + archive_name + " ...") execute_tar(tmp_dir, "-czf " + File.bash_path(release.dist_dir + Path.explode(archive_name)) + " " + Bash.string(isabelle_name)) case Platform.Family.macos => File.change(isabelle_target + jedit_props, _.replaceAll("delete-line.shortcut=.*", "delete-line.shortcut=C+d") .replaceAll("delete.shortcut2=.*", "delete.shortcut2=A+d")) // macOS application bundle val app_contents = isabelle_target + Path.explode("Contents") for (icon <- List("lib/logo/isabelle.icns", "lib/logo/theory.icns")) { Isabelle_System.copy_file(isabelle_target + Path.explode(icon), Isabelle_System.make_directory(app_contents + Path.explode("Resources"))) } make_isabelle_plist( app_contents + Path.explode("Info.plist"), isabelle_name, release.ident) make_isabelle_app(platform, isabelle_target, isabelle_name, jdk_component, classpath, dock_icon = true) val isabelle_options = Path.explode("Isabelle.options") make_isabelle_options( isabelle_target + isabelle_options, java_options ::: List("-Disabelle.app=true")) // application archive val archive_name = isabelle_name + "_macos.tar.gz" progress.echo("Packaging " + archive_name + " ...") val isabelle_app = Path.explode(isabelle_name + ".app") Isabelle_System.move_file(tmp_dir + Path.explode(isabelle_name), tmp_dir + isabelle_app) execute_tar(tmp_dir, "-czf " + File.bash_path(release.dist_dir + Path.explode(archive_name)) + " " + File.bash_path(isabelle_app)) case Platform.Family.windows => File.change(isabelle_target + jedit_props, _.replaceAll("foldPainter=.*", "foldPainter=Square")) // application launcher Isabelle_System.move_file(isabelle_target + Path.explode("contrib/windows_app"), tmp_dir) val app_template = Path.explode("~~/Admin/Windows/launch4j") make_isabelle_options( isabelle_target + Path.explode(isabelle_name + ".l4j.ini"), java_options, line_ending = "\r\n") val isabelle_xml = Path.explode("isabelle.xml") val isabelle_exe = Path.explode(isabelle_name + ".exe") File.write(tmp_dir + isabelle_xml, File.read(app_template + isabelle_xml) .replace("{ISABELLE_NAME}", isabelle_name) .replace("{OUTFILE}", File.platform_path(isabelle_target + isabelle_exe)) .replace("{ICON}", File.platform_path(app_template + Path.explode("isabelle_transparent.ico"))) .replace("{SPLASH}", File.platform_path(app_template + Path.explode("isabelle.bmp"))) .replace("{CLASSPATH}", cat_lines(classpath.map(cp => " %EXEDIR%\\" + File.platform_path(cp).replace('/', '\\') + ""))) .replace("\\jdk\\", "\\" + jdk_component + "\\")) execute(tmp_dir, "\"windows_app/launch4j-${ISABELLE_PLATFORM_FAMILY}/launch4j\" isabelle.xml") Isabelle_System.copy_file(app_template + Path.explode("manifest.xml"), isabelle_target + isabelle_exe.ext("manifest")) // Cygwin setup val cygwin_template = Path.explode("~~/Admin/Windows/Cygwin") Isabelle_System.copy_file(cygwin_template + Path.explode("Cygwin-Terminal.bat"), isabelle_target) val cygwin_mirror = File.read(isabelle_target + Path.explode("contrib/cygwin/isabelle/cygwin_mirror")) val cygwin_bat = Path.explode("Cygwin-Setup.bat") File.write(isabelle_target + cygwin_bat, File.read(cygwin_template + cygwin_bat).replace("{MIRROR}", cygwin_mirror)) File.set_executable(isabelle_target + cygwin_bat, true) for (name <- List("isabelle/postinstall", "isabelle/rebaseall")) { val path = Path.explode(name) Isabelle_System.copy_file(cygwin_template + path, isabelle_target + Path.explode("contrib/cygwin") + path) } execute(isabelle_target, """find . -type f -not -name "*.exe" -not -name "*.dll" """ + (if (Platform.is_macos) "-perm +100" else "-executable") + " -print0 > contrib/cygwin/isabelle/executables") execute(isabelle_target, """find . -type l -exec echo "{}" ";" -exec readlink "{}" ";" """ + """> contrib/cygwin/isabelle/symlinks""") execute(isabelle_target, """find . -type l -exec rm "{}" ";" """) File.write(isabelle_target + Path.explode("contrib/cygwin/isabelle/uninitialized"), "") // executable archive (self-extracting 7z) val archive_name = isabelle_name + ".7z" val exe_archive = tmp_dir + Path.explode(archive_name) exe_archive.file.delete progress.echo("Packaging " + archive_name + " ...") execute(tmp_dir, "7z -y -bd a " + File.bash_path(exe_archive) + " " + Bash.string(isabelle_name)) if (!exe_archive.is_file) error("Failed to create archive: " + exe_archive) val sfx_exe = tmp_dir + Path.explode("windows_app/7zsd_All_x64.sfx") val sfx_txt = File.read(Path.explode("~~/Admin/Windows/Installer/sfx.txt")) .replace("{ISABELLE_NAME}", isabelle_name) Bytes.write(release.dist_dir + isabelle_exe, Bytes.read(sfx_exe) + Bytes(sfx_txt) + Bytes.read(exe_archive)) File.set_executable(release.dist_dir + isabelle_exe, true) } }) progress.echo("DONE") } /* minimal website */ for (dir <- website) { val website_platform_bundles = for { bundle_info <- bundle_infos if (release.dist_dir + bundle_info.path).is_file } yield (bundle_info.name, bundle_info) val isabelle_link = HTML.link(Isabelle_Cronjob.isabelle_repos_source + "/rev/" + release.ident, HTML.text("Isabelle/" + release.ident)) val afp_link = HTML.link(AFP.repos_source + "/rev/" + afp_rev, HTML.text("AFP/" + afp_rev)) HTML.write_document(dir, "index.html", List(HTML.title(release.dist_name)), List( HTML.section(release.dist_name), HTML.subsection("Platforms"), HTML.itemize( website_platform_bundles.map({ case (bundle, bundle_info) => List(HTML.link(bundle, HTML.text(bundle_info.platform_description))) })), HTML.subsection("Repositories"), HTML.itemize( List(List(isabelle_link)) ::: (if (afp_rev == "") Nil else List(List(afp_link)))))) for ((bundle, _) <- website_platform_bundles) Isabelle_System.copy_file(release.dist_dir + Path.explode(bundle), dir) } /* HTML library */ if (build_library) { if (release.isabelle_library_archive.is_file) { progress.echo_warning("Library archive already exists: " + release.isabelle_library_archive) } else { Isabelle_System.with_tmp_dir("build_release")(tmp_dir => { val bundle = release.dist_dir + Path.explode(release.dist_name + "_" + Platform.family + ".tar.gz") execute_tar(tmp_dir, "-xzf " + File.bash_path(bundle)) val other_isabelle = release.other_isabelle(tmp_dir) Isabelle_System.make_directory(other_isabelle.etc) File.write(other_isabelle.etc_preferences, "ML_system_64 = true\n") other_isabelle.bash("bin/isabelle build -f -j " + parallel_jobs + " -o browser_info -o document=pdf -o document_variants=document:outline=/proof,/ML" + " -o system_heaps -c -a -d '~~/src/Benchmarks'", echo = true).check other_isabelle.isabelle_home_user.file.delete execute(tmp_dir, "chmod -R a+r " + Bash.string(release.dist_name)) execute(tmp_dir, "chmod -R g=o " + Bash.string(release.dist_name)) execute_tar(tmp_dir, "-czf " + File.bash_path(release.isabelle_library_archive) + " " + Bash.string(release.dist_name + "/browser_info")) }) } } release } /** command line entry point **/ def main(args: Array[String]): Unit = { Command_Line.tool { var afp_rev = "" var components_base: Path = Components.default_components_base var proper_release_name: Option[String] = None var website: Option[Path] = None var build_sessions: List[String] = Nil var more_components: List[Path] = Nil var parallel_jobs = 1 var build_library = false var options = Options.init() var platform_families = default_platform_families var rev = "" val getopts = Getopts(""" Usage: Admin/build_release [OPTIONS] BASE_DIR Options are: -A REV corresponding AFP changeset id -C DIR base directory for Isabelle components (default: """ + Components.default_components_base + """) -R RELEASE proper release with name -W WEBSITE produce minimal website in given directory -b SESSIONS build platform-specific session images (separated by commas) -c ARCHIVE clean bundling with additional component .tar.gz archive -j INT maximum number of parallel jobs (default 1) -l build library -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NAMES platform families (default: """ + default_platform_families.mkString(",") + """) -r REV Mercurial changeset id (default: RELEASE or tip) Build Isabelle release in base directory, using the local repository clone. """, "A:" -> (arg => afp_rev = arg), "C:" -> (arg => components_base = Path.explode(arg)), "R:" -> (arg => proper_release_name = Some(arg)), "W:" -> (arg => website = Some(Path.explode(arg))), "b:" -> (arg => build_sessions = space_explode(',', arg)), "c:" -> (arg => { val path = Path.explode(arg) Components.Archive.get_name(path.file_name) more_components = more_components ::: List(path) }), "j:" -> (arg => parallel_jobs = Value.Int.parse(arg)), "l" -> (_ => build_library = true), "o:" -> (arg => options = options + arg), "p:" -> (arg => platform_families = space_explode(',', arg).map(Platform.Family.parse)), "r:" -> (arg => rev = arg)) val more_args = getopts(args) val base_dir = more_args match { case List(base_dir) => base_dir case _ => getopts.usage() } val progress = new Console_Progress() if (platform_families.contains(Platform.Family.windows) && !Isabelle_System.bash("7z i").ok) error("Building for windows requires 7z") build_release(Path.explode(base_dir), options, components_base = components_base, progress = progress, rev = rev, afp_rev = afp_rev, proper_release_name = proper_release_name, website = website, platform_families = if (platform_families.isEmpty) default_platform_families else platform_families, more_components = more_components, build_sessions = build_sessions, build_library = build_library, parallel_jobs = parallel_jobs) } } } 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,305 @@ /* 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 + ": " + Distribution.version)) + Some(Response.text("Welcome to Isabelle/" + id + Isabelle_System.isabelle_heading())) } 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/PIDE/session.ML b/src/Pure/PIDE/session.ML --- a/src/Pure/PIDE/session.ML +++ b/src/Pure/PIDE/session.ML @@ -1,62 +1,59 @@ (* Title: Pure/PIDE/session.ML Author: Makarius Prover session: persistent state of logic image. *) signature SESSION = sig val init: string -> unit val get_name: unit -> string val welcome: unit -> string val get_keywords: unit -> Keyword.keywords val shutdown: unit -> unit val finish: unit -> unit end; structure Session: SESSION = struct (* session name *) val session = Synchronized.var "Session.session" ""; fun init name = Synchronized.change session (K name); fun get_name () = Synchronized.value session; fun description () = "Isabelle/" ^ get_name (); -fun welcome () = - if Distribution.is_identified then - "Welcome to " ^ description () ^ " (" ^ Distribution.version ^ ")" - else "Unofficial version of " ^ description () ^ " (" ^ Distribution.version ^ ")"; +fun welcome () = "Welcome to " ^ description () ^ Isabelle_System.isabelle_heading (); (* base syntax *) val keywords = Synchronized.var "Session.keywords" Keyword.empty_keywords; fun get_keywords () = Synchronized.value keywords; fun update_keywords () = Synchronized.change keywords (K (fold (curry Keyword.merge_keywords o Thy_Header.get_keywords o Thy_Info.get_theory) (Thy_Info.get_names ()) Keyword.empty_keywords)); (* finish *) fun shutdown () = (Execution.shutdown (); Event_Timer.shutdown (); Future.shutdown ()); fun finish () = (shutdown (); Par_List.map (Global_Theory.get_thm_names o Thy_Info.get_theory) (Thy_Info.get_names ()); Thy_Info.finish (); shutdown (); update_keywords ()); end; diff --git a/src/Pure/ROOT.ML b/src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML +++ b/src/Pure/ROOT.ML @@ -1,358 +1,357 @@ (* Title: Pure/ROOT.ML Author: Makarius Main entry point for the Isabelle/Pure bootstrap process. Note: When this file is open in the Prover IDE, the ML files of Isabelle/Pure can be explored interactively. This is a separate copy of Pure within Pure: it does not affect the running logic session. *) chapter "Isabelle/Pure bootstrap"; ML_file "ML/ml_name_space.ML"; section "Bootstrap phase 0: Poly/ML setup"; ML_file "ML/ml_init.ML"; ML_file "ML/ml_system.ML"; -ML_file "System/distribution.ML"; ML_file "General/basics.ML"; ML_file "General/symbol_explode.ML"; ML_file "Concurrent/multithreading.ML"; ML_file "Concurrent/unsynchronized.ML"; ML_file "Concurrent/synchronized.ML"; ML_file "Concurrent/counter.ML"; ML_file "ML/ml_heap.ML"; ML_file "ML/ml_profiling.ML"; ML_file "ML/ml_print_depth0.ML"; ML_file "ML/ml_pretty.ML"; ML_file "ML/ml_compiler0.ML"; section "Bootstrap phase 1: towards ML within position context"; subsection "Library of general tools"; ML_file "library.ML"; ML_file "General/print_mode.ML"; ML_file "General/alist.ML"; ML_file "General/table.ML"; ML_file "General/random.ML"; ML_file "General/value.ML"; ML_file "General/properties.ML"; ML_file "General/output.ML"; ML_file "PIDE/markup.ML"; ML_file "General/utf8.ML"; ML_file "General/scan.ML"; ML_file "General/source.ML"; ML_file "General/symbol.ML"; ML_file "General/position.ML"; ML_file "General/symbol_pos.ML"; ML_file "General/input.ML"; ML_file "General/comment.ML"; ML_file "General/antiquote.ML"; ML_file "ML/ml_lex.ML"; ML_file "ML/ml_compiler1.ML"; section "Bootstrap phase 2: towards ML within Isar context"; subsection "Library of general tools"; ML_file "General/integer.ML"; ML_file "General/stack.ML"; ML_file "General/queue.ML"; ML_file "General/heap.ML"; ML_file "General/same.ML"; ML_file "General/ord_list.ML"; ML_file "General/balanced_tree.ML"; ML_file "General/linear_set.ML"; ML_file "General/buffer.ML"; ML_file "General/pretty.ML"; ML_file "General/rat.ML"; ML_file "PIDE/xml.ML"; ML_file "General/path.ML"; ML_file "General/url.ML"; ML_file "System/bash.ML"; ML_file "General/file.ML"; ML_file "General/long_name.ML"; ML_file "General/binding.ML"; ML_file "General/socket_io.ML"; ML_file "General/seq.ML"; ML_file "General/time.ML"; ML_file "General/timing.ML"; ML_file "General/sha1.ML"; ML_file "PIDE/byte_message.ML"; ML_file "PIDE/yxml.ML"; ML_file "PIDE/protocol_message.ML"; ML_file "PIDE/document_id.ML"; ML_file "General/change_table.ML"; ML_file "General/graph.ML"; ML_file "System/options.ML"; subsection "Fundamental structures"; ML_file "name.ML"; ML_file "term.ML"; ML_file "context.ML"; ML_file "config.ML"; ML_file "context_position.ML"; ML_file "soft_type_system.ML"; subsection "Concurrency within the ML runtime"; ML_file "ML/exn_properties.ML"; ML_file_no_debug "ML/exn_debugger.ML"; ML_file "Concurrent/thread_data_virtual.ML"; ML_file "Concurrent/isabelle_thread.ML"; ML_file "Concurrent/single_assignment.ML"; ML_file "Concurrent/par_exn.ML"; ML_file "Concurrent/task_queue.ML"; ML_file "Concurrent/future.ML"; ML_file "Concurrent/event_timer.ML"; ML_file "Concurrent/timeout.ML"; ML_file "Concurrent/lazy.ML"; ML_file "Concurrent/par_list.ML"; ML_file "Concurrent/mailbox.ML"; ML_file "Concurrent/cache.ML"; ML_file "PIDE/active.ML"; ML_file "Thy/export.ML"; subsection "Inner syntax"; ML_file "Syntax/type_annotation.ML"; ML_file "Syntax/term_position.ML"; ML_file "Syntax/lexicon.ML"; ML_file "Syntax/ast.ML"; ML_file "Syntax/syntax_ext.ML"; ML_file "Syntax/parser.ML"; ML_file "Syntax/syntax_trans.ML"; ML_file "Syntax/mixfix.ML"; ML_file "Syntax/printer.ML"; ML_file "Syntax/syntax.ML"; subsection "Core of tactical proof system"; ML_file "term_ord.ML"; ML_file "term_subst.ML"; ML_file "General/completion.ML"; ML_file "General/name_space.ML"; ML_file "sorts.ML"; ML_file "type.ML"; ML_file "logic.ML"; ML_file "Syntax/simple_syntax.ML"; ML_file "net.ML"; ML_file "item_net.ML"; ML_file "envir.ML"; ML_file "consts.ML"; ML_file "term_xml.ML"; ML_file "primitive_defs.ML"; ML_file "sign.ML"; ML_file "defs.ML"; ML_file "term_sharing.ML"; ML_file "pattern.ML"; ML_file "unify.ML"; ML_file "theory.ML"; ML_file "proofterm.ML"; ML_file "thm.ML"; ML_file "more_pattern.ML"; ML_file "more_unify.ML"; ML_file "more_thm.ML"; ML_file "facts.ML"; ML_file "thm_name.ML"; ML_file "global_theory.ML"; ML_file "pure_thy.ML"; ML_file "drule.ML"; ML_file "morphism.ML"; ML_file "variable.ML"; ML_file "conv.ML"; ML_file "goal_display.ML"; ML_file "tactical.ML"; ML_file "search.ML"; ML_file "tactic.ML"; ML_file "raw_simplifier.ML"; ML_file "conjunction.ML"; ML_file "assumption.ML"; subsection "Isar -- Intelligible Semi-Automated Reasoning"; (*ML support and global execution*) ML_file "ML/ml_syntax.ML"; ML_file "ML/ml_env.ML"; ML_file "ML/ml_options.ML"; ML_file "ML/ml_print_depth.ML"; ML_file_no_debug "Isar/runtime.ML"; ML_file "PIDE/execution.ML"; ML_file "ML/ml_compiler.ML"; ML_file "skip_proof.ML"; ML_file "goal.ML"; (*outer syntax*) ML_file "Isar/keyword.ML"; ML_file "Isar/token.ML"; ML_file "Isar/parse.ML"; ML_file "Thy/document_source.ML"; ML_file "Thy/thy_header.ML"; ML_file "Thy/document_marker.ML"; (*proof context*) ML_file "Isar/object_logic.ML"; ML_file "Isar/rule_cases.ML"; ML_file "Isar/auto_bind.ML"; ML_file "type_infer.ML"; ML_file "Syntax/local_syntax.ML"; ML_file "Isar/proof_context.ML"; ML_file "type_infer_context.ML"; ML_file "Syntax/syntax_phases.ML"; ML_file "Isar/args.ML"; (*theory specifications*) ML_file "Isar/local_defs.ML"; ML_file "Isar/local_theory.ML"; ML_file "Isar/entity.ML"; ML_file "PIDE/command_span.ML"; ML_file "Thy/thy_element.ML"; ML_file "Thy/markdown.ML"; ML_file "Thy/latex.ML"; (*ML with context and antiquotations*) ML_file "ML/ml_context.ML"; ML_file "ML/ml_antiquotation.ML"; ML_file "ML/ml_compiler2.ML"; ML_file "ML/ml_pid.ML"; section "Bootstrap phase 3: towards theory Pure and final ML toplevel setup"; (*basic proof engine*) ML_file "par_tactical.ML"; ML_file "context_tactic.ML"; ML_file "Isar/proof_display.ML"; ML_file "Isar/attrib.ML"; ML_file "Isar/context_rules.ML"; ML_file "Isar/method.ML"; ML_file "Isar/proof.ML"; ML_file "Isar/element.ML"; ML_file "Isar/obtain.ML"; ML_file "Isar/subgoal.ML"; ML_file "Isar/calculation.ML"; (*local theories and targets*) ML_file "Isar/locale.ML"; ML_file "Isar/generic_target.ML"; ML_file "Isar/bundle.ML"; ML_file "Isar/overloading.ML"; ML_file "axclass.ML"; ML_file "Isar/class.ML"; ML_file "Isar/named_target.ML"; ML_file "Isar/expression.ML"; ML_file "Isar/interpretation.ML"; ML_file "Isar/class_declaration.ML"; ML_file "Isar/target_context.ML"; ML_file "Isar/experiment.ML"; ML_file "simplifier.ML"; ML_file "Tools/plugin.ML"; (*executable theory content*) ML_file "Isar/code.ML"; (*specifications*) ML_file "Isar/spec_rules.ML"; ML_file "Isar/specification.ML"; ML_file "Isar/parse_spec.ML"; ML_file "Isar/typedecl.ML"; (*toplevel transactions*) ML_file "Isar/proof_node.ML"; ML_file "Isar/toplevel.ML"; (*proof term operations*) ML_file "Proof/proof_rewrite_rules.ML"; ML_file "Proof/proof_syntax.ML"; ML_file "Proof/proof_checker.ML"; ML_file "Proof/extraction.ML"; (*Isabelle system*) ML_file "PIDE/protocol_command.ML"; ML_file "System/scala.ML"; ML_file "System/process_result.ML"; ML_file "System/isabelle_system.ML"; (*theory documents*) ML_file "Thy/term_style.ML"; ML_file "Isar/outer_syntax.ML"; ML_file "ML/ml_antiquotations.ML"; ML_file "Thy/document_antiquotation.ML"; ML_file "Thy/thy_output.ML"; ML_file "Thy/document_antiquotations.ML"; ML_file "General/graph_display.ML"; ML_file "pure_syn.ML"; ML_file "PIDE/command.ML"; ML_file "PIDE/query_operation.ML"; ML_file "PIDE/resources.ML"; ML_file "Thy/thy_info.ML"; ML_file "thm_deps.ML"; ML_file "Thy/export_theory.ML"; ML_file "Thy/sessions.ML"; ML_file "PIDE/session.ML"; ML_file "PIDE/document.ML"; (*theory and proof operations*) ML_file "Isar/isar_cmd.ML"; subsection "Isabelle/Isar system"; ML_file "System/command_line.ML"; ML_file "System/message_channel.ML"; ML_file "System/isabelle_process.ML"; ML_file "System/scala_compiler.ML"; ML_file "System/isabelle_tool.ML"; ML_file "Thy/bibtex.ML"; ML_file "PIDE/protocol.ML"; ML_file "General/output_primitives_virtual.ML"; subsection "Miscellaneous tools and packages for Pure Isabelle"; ML_file "ML/ml_pp.ML"; ML_file "ML/ml_thms.ML"; ML_file "ML/ml_file.ML"; ML_file "Tools/build.ML"; ML_file "Tools/named_thms.ML"; ML_file "Tools/print_operation.ML"; ML_file "Tools/rail.ML"; ML_file "Tools/rule_insts.ML"; ML_file "Tools/thy_deps.ML"; ML_file "Tools/class_deps.ML"; ML_file "Tools/find_theorems.ML"; ML_file "Tools/find_consts.ML"; ML_file "Tools/simplifier_trace.ML"; ML_file_no_debug "Tools/debugger.ML"; ML_file "Tools/named_theorems.ML"; ML_file "Tools/doc.ML"; ML_file "Tools/jedit.ML"; ML_file "Tools/ghc.ML"; ML_file "Tools/generated_files.ML" diff --git a/src/Pure/System/distribution.ML b/src/Pure/System/distribution.ML deleted file mode 100644 --- a/src/Pure/System/distribution.ML +++ /dev/null @@ -1,11 +0,0 @@ -(* Title: Pure/System/distribution.ML - Author: Makarius - -The Isabelle system distribution -- filled-in by makedist. -*) - -structure Distribution = -struct - val version = "repository version"; - val is_identified = false; -end; diff --git a/src/Pure/System/distribution.scala b/src/Pure/System/distribution.scala deleted file mode 100644 --- a/src/Pure/System/distribution.scala +++ /dev/null @@ -1,14 +0,0 @@ -/* Title: Pure/System/distribution.scala - Author: Makarius - -The Isabelle system distribution -- filled-in by makedist. -*/ - -package isabelle - - -object Distribution -{ - val version = "repository version" - val is_identified = false -} 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,118 +1,136 @@ (* 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 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"; + 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,597 +1,613 @@ /* 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): Option[String] = getetc("ISABELLE_ID", root = root) orElse Mercurial.archive_id(root) orElse { if (Mercurial.is_repository(root)) Some(Mercurial.repository(root).parent()) else None } + object Isabelle_Id extends Scala.Fun("isabelle_id") + { + val here = Scala_Project.here + def apply(arg: String): String = isabelle_id().get + } + 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") + /** 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)); "" } } } diff --git a/src/Pure/System/scala.scala b/src/Pure/System/scala.scala --- a/src/Pure/System/scala.scala +++ b/src/Pure/System/scala.scala @@ -1,264 +1,265 @@ /* Title: Pure/System/scala.scala Author: Makarius Support for Scala at runtime. */ package isabelle import java.io.{File => JFile, StringWriter, PrintWriter} import scala.tools.nsc.{GenericRunnerSettings, ConsoleWriter, NewLinePrintWriter} import scala.tools.nsc.interpreter.{IMain, Results} import scala.tools.nsc.interpreter.shell.ReplReporterImpl object Scala { /** registered functions **/ abstract class Fun(val name: String, val thread: Boolean = false) extends Function[String, String] { override def toString: String = name def position: Properties.T = here.position def here: Scala_Project.Here def apply(arg: String): String } class Functions(val functions: Fun*) extends Isabelle_System.Service lazy val functions: List[Fun] = Isabelle_System.make_services(classOf[Functions]).flatMap(_.functions) /** demo functions **/ object Echo extends Fun("echo") { val here = Scala_Project.here def apply(arg: String): String = arg } object Sleep extends Fun("sleep") { val here = Scala_Project.here def apply(seconds: String): String = { val t = seconds match { case Value.Double(s) => Time.seconds(s) case _ => error("Malformed argument: " + quote(seconds)) } val t0 = Time.now() t.sleep val t1 = Time.now() (t1 - t0).toString } } /** compiler **/ object Compiler { def context( error: String => Unit = Exn.error, jar_dirs: List[JFile] = Nil): Context = { def find_jars(dir: JFile): List[String] = File.find_files(dir, file => file.getName.endsWith(".jar")). map(File.absolute_name) val class_path = for { prop <- List("isabelle.scala.classpath", "java.class.path") path = System.getProperty(prop, "") if path != "\"\"" elem <- space_explode(JFile.pathSeparatorChar, path) } yield elem val settings = new GenericRunnerSettings(error) settings.classpath.value = (class_path ::: jar_dirs.flatMap(find_jars)).mkString(JFile.pathSeparator) new Context(settings) } def default_print_writer: PrintWriter = new NewLinePrintWriter(new ConsoleWriter, true) class Context private [Compiler](val settings: GenericRunnerSettings) { override def toString: String = settings.toString def interpreter( print_writer: PrintWriter = default_print_writer, class_loader: ClassLoader = null): IMain = { new IMain(settings, new ReplReporterImpl(settings, print_writer)) { override def parentClassLoader: ClassLoader = if (class_loader == null) super.parentClassLoader else class_loader } } def toplevel(interpret: Boolean, source: String): List[String] = { val out = new StringWriter val interp = interpreter(new PrintWriter(out)) val marker = '\u000b' val ok = interp.withLabel(marker.toString) { if (interpret) interp.interpret(source) == Results.Success else (new interp.ReadEvalPrint).compile(source) } out.close() val Error = """(?s)^\S* error: (.*)$""".r val errors = space_explode(marker, Library.strip_ansi_color(out.toString)). collect({ case Error(msg) => "Scala error: " + Library.trim_line(msg) }) if (!ok && errors.isEmpty) List("Error") else errors } } } object Toplevel extends Fun("scala_toplevel") { val here = Scala_Project.here def apply(arg: String): String = { val (interpret, source) = YXML.parse_body(arg) match { case Nil => (false, "") case List(XML.Text(source)) => (false, source) case body => import XML.Decode._; pair(bool, string)(body) } val errors = try { Compiler.context().toplevel(interpret, source) } catch { case ERROR(msg) => List(msg) } locally { import XML.Encode._; YXML.string_of_body(list(string)(errors)) } } } /** invoke Scala functions from ML **/ /* invoke function */ object Tag extends Enumeration { val NULL, OK, ERROR, FAIL, INTERRUPT = Value } def function_thread(name: String): Boolean = functions.find(fun => fun.name == name) match { case Some(fun) => fun.thread case None => false } def function_body(name: String, arg: String): (Tag.Value, String) = functions.find(fun => fun.name == name) match { case Some(fun) => Exn.capture { fun(arg) } match { case Exn.Res(null) => (Tag.NULL, "") case Exn.Res(res) => (Tag.OK, res) case Exn.Exn(Exn.Interrupt()) => (Tag.INTERRUPT, "") case Exn.Exn(e) => (Tag.ERROR, Exn.message(e)) } case None => (Tag.FAIL, "Unknown Isabelle/Scala function: " + quote(name)) } /* protocol handler */ class Handler extends Session.Protocol_Handler { private var session: Session = null private var futures = Map.empty[String, Future[Unit]] override def init(session: Session): Unit = synchronized { this.session = session } override def exit(): Unit = synchronized { for ((id, future) <- futures) cancel(id, future) futures = Map.empty } private def result(id: String, tag: Scala.Tag.Value, res: String): Unit = synchronized { if (futures.isDefinedAt(id)) { session.protocol_command("Scala.result", id, tag.id.toString, res) futures -= id } } private def cancel(id: String, future: Future[Unit]): Unit = { future.cancel() result(id, Scala.Tag.INTERRUPT, "") } private def invoke_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Invoke_Scala(name, id) => def body: Unit = { val (tag, res) = Scala.function_body(name, msg.text) result(id, tag, res) } val future = if (Scala.function_thread(name)) { Future.thread(name = Isabelle_Thread.make_name(base = "invoke_scala"))(body) } else Future.fork(body) futures += (id -> future) true case _ => false } } private def cancel_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Cancel_Scala(id) => futures.get(id) match { case Some(future) => cancel(id, future) case None => } true case _ => false } } override val functions = List( Markup.Invoke_Scala.name -> invoke_scala, Markup.Cancel_Scala.name -> cancel_scala) } } class Scala_Functions extends Scala.Functions( Scala.Echo, Scala.Sleep, Scala.Toplevel, Doc.Doc_Names, Bash.Process, Bibtex.Check_Database, Isabelle_System.Make_Directory, Isabelle_System.Copy_Dir, Isabelle_System.Copy_File, Isabelle_System.Copy_File_Base, Isabelle_System.Rm_Tree, Isabelle_System.Download, + Isabelle_System.Isabelle_Id, Isabelle_Tool.Isabelle_Tools, isabelle.atp.SystemOnTPTP.List_Systems, isabelle.atp.SystemOnTPTP.Run_System) diff --git a/src/Pure/Thy/presentation.scala b/src/Pure/Thy/presentation.scala --- a/src/Pure/Thy/presentation.scala +++ b/src/Pure/Thy/presentation.scala @@ -1,793 +1,793 @@ /* Title: Pure/Thy/present.scala Author: Makarius HTML/PDF presentation of theory documents. */ package isabelle import scala.collection.immutable.SortedMap object Presentation { /** HTML documents **/ val fonts_path = Path.explode("fonts") sealed case class HTML_Document(title: String, content: String) def html_context(fonts_url: String => String = HTML.fonts_url()): HTML_Context = new HTML_Context(fonts_url) final class HTML_Context private[Presentation](fonts_url: String => String) { def init_fonts(dir: Path): Unit = { val fonts_dir = Isabelle_System.make_directory(dir + fonts_path) for (entry <- Isabelle_Fonts.fonts(hidden = true)) Isabelle_System.copy_file(entry.path, fonts_dir) } def head(title: String, rest: XML.Body = Nil): XML.Tree = HTML.div("head", HTML.chapter(title) :: rest) def source(body: XML.Body): XML.Tree = HTML.pre("source", body) def contents(heading: String, items: List[XML.Body], css_class: String = "contents") : List[XML.Elem] = { if (items.isEmpty) Nil else List(HTML.div(css_class, List(HTML.section(heading), HTML.itemize(items)))) } def output_document(title: String, body: XML.Body): String = HTML.output_document( List( HTML.style(HTML.fonts_css(fonts_url) + "\n\n" + File.read(HTML.isabelle_css)), HTML.title(title)), List(HTML.source(body)), css = "", structural = false) def html_document(title: String, body: XML.Body): HTML_Document = HTML_Document(title, output_document(title, body)) } /* presentation elements */ sealed case class Elements( html: Markup.Elements = Markup.Elements.empty, entity: Markup.Elements = Markup.Elements.empty, language: Markup.Elements = Markup.Elements.empty) val elements1: Elements = Elements( html = Rendering.foreground_elements ++ Rendering.text_color_elements + Markup.NUMERAL + Markup.COMMENT + Markup.ENTITY + Markup.LANGUAGE, entity = Markup.Elements(Markup.THEORY)) val elements2: Elements = Elements( html = elements1.html ++ Rendering.markdown_elements, language = Markup.Elements(Markup.Language.DOCUMENT)) /* HTML */ private val div_elements = Set(HTML.div.name, HTML.pre.name, HTML.par.name, HTML.list.name, HTML.enum.name, HTML.descr.name) def make_html( elements: Elements, entity_link: (Properties.T, XML.Body) => Option[XML.Tree], xml: XML.Body): XML.Body = { def html_div(html: XML.Body): Boolean = html exists { case XML.Elem(markup, body) => div_elements.contains(markup.name) || html_div(body) case XML.Text(_) => false } def html_class(c: String, html: XML.Body): XML.Body = if (c == "") html else if (html_div(html)) List(HTML.div(c, html)) else List(HTML.span(c, html)) def html_body(xml_body: XML.Body): XML.Body = xml_body flatMap { case XML.Wrapped_Elem(markup, _, body) => html_body(List(XML.Elem(markup, body))) case XML.Elem(Markup(Markup.ENTITY, props @ Markup.Kind(kind)), body) => val body1 = html_body(body) if (elements.entity(kind)) { entity_link(props, body1) match { case Some(link) => List(link) case None => body1 } } else body1 case XML.Elem(Markup(Markup.LANGUAGE, Markup.Name(name)), body) => html_class(if (elements.language(name)) name else "", html_body(body)) case XML.Elem(Markup(Markup.MARKDOWN_PARAGRAPH, _), body) => List(HTML.par(html_body(body))) case XML.Elem(Markup(Markup.MARKDOWN_ITEM, _), body) => List(HTML.item(html_body(body))) case XML.Elem(Markup(Markup.Markdown_Bullet.name, _), _) => Nil case XML.Elem(Markup.Markdown_List(kind), body) => if (kind == Markup.ENUMERATE) List(HTML.enum(html_body(body))) else List(HTML.list(html_body(body))) case XML.Elem(markup, body) => val name = markup.name val html = markup.properties match { case Markup.Kind(kind) if kind == Markup.COMMAND || kind == Markup.KEYWORD => html_class(kind, html_body(body)) case _ => html_body(body) } Rendering.foreground.get(name) orElse Rendering.text_color.get(name) match { case Some(c) => html_class(c.toString, html) case None => html_class(name, html) } case XML.Text(text) => HTML.text(Symbol.decode(text)) } html_body(xml) } /* PIDE HTML document */ def html_document( resources: Resources, snapshot: Document.Snapshot, html_context: HTML_Context, elements: Elements, plain_text: Boolean = false): HTML_Document = { require(!snapshot.is_outdated, "document snapshot outdated") val name = snapshot.node_name if (plain_text) { val title = "File " + Symbol.cartouche_decoded(name.path.file_name) val body = HTML.text(snapshot.node.source) html_context.html_document(title, body) } else { resources.html_document(snapshot) getOrElse { val title = if (name.is_theory) "Theory " + quote(name.theory_base_name) else "File " + Symbol.cartouche_decoded(name.path.file_name) val body = make_html(elements, (_, _) => None, snapshot.xml_markup(elements = elements.html)) html_context.html_document(title, body) } } } /** PDF LaTeX documents **/ /* document info */ abstract class Document_Name { def name: String def path: Path = Path.basic(name) override def toString: String = name } object Document_Variant { def parse(name: String, tags: String): Document_Variant = Document_Variant(name, Library.space_explode(',', tags)) def parse(opt: String): Document_Variant = Library.space_explode('=', opt) match { case List(name) => Document_Variant(name, Nil) case List(name, tags) => parse(name, tags) case _ => error("Malformed document variant: " + quote(opt)) } } sealed case class Document_Variant(name: String, tags: List[String]) extends Document_Name { def print_tags: String = tags.mkString(",") def print: String = if (tags.isEmpty) name else name + "=" + print_tags def latex_sty: String = Library.terminate_lines( tags.map(tag => tag.toList match { case '/' :: cs => "\\isafoldtag{" + cs.mkString + "}" case '-' :: cs => "\\isadroptag{" + cs.mkString + "}" case '+' :: cs => "\\isakeeptag{" + cs.mkString + "}" case cs => "\\isakeeptag{" + cs.mkString + "}" })) } sealed case class Document_Input(name: String, sources: SHA1.Digest) extends Document_Name sealed case class Document_Output(name: String, sources: SHA1.Digest, log_xz: Bytes, pdf: Bytes) extends Document_Name { def log: String = log_xz.uncompress().text def log_lines: List[String] = split_lines(log) def write(db: SQL.Database, session_name: String): Unit = write_document(db, session_name, this) } /* SQL data model */ object Data { val session_name = SQL.Column.string("session_name").make_primary_key val name = SQL.Column.string("name").make_primary_key val sources = SQL.Column.string("sources") val log_xz = SQL.Column.bytes("log_xz") val pdf = SQL.Column.bytes("pdf") val table = SQL.Table("isabelle_documents", List(session_name, name, sources, log_xz, pdf)) def where_equal(session_name: String, name: String = ""): SQL.Source = "WHERE " + Data.session_name.equal(session_name) + (if (name == "") "" else " AND " + Data.name.equal(name)) } def read_documents(db: SQL.Database, session_name: String): List[Document_Input] = { val select = Data.table.select(List(Data.name, Data.sources), Data.where_equal(session_name)) db.using_statement(select)(stmt => stmt.execute_query().iterator(res => { val name = res.string(Data.name) val sources = res.string(Data.sources) Document_Input(name, SHA1.fake(sources)) }).toList) } def read_document(db: SQL.Database, session_name: String, name: String): Option[Document_Output] = { val select = Data.table.select(sql = Data.where_equal(session_name, name)) db.using_statement(select)(stmt => { val res = stmt.execute_query() if (res.next()) { val name = res.string(Data.name) val sources = res.string(Data.sources) val log_xz = res.bytes(Data.log_xz) val pdf = res.bytes(Data.pdf) Some(Document_Output(name, SHA1.fake(sources), log_xz, pdf)) } else None }) } def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit = { db.using_statement(Data.table.insert())(stmt => { stmt.string(1) = session_name stmt.string(2) = doc.name stmt.string(3) = doc.sources.toString stmt.bytes(4) = doc.log_xz stmt.bytes(5) = doc.pdf stmt.execute() }) } /* presentation context */ object Context { val none: Context = new Context { def enabled: Boolean = false } val standard: Context = new Context { def enabled: Boolean = true } def dir(path: Path): Context = new Context { def enabled: Boolean = true override def dir(store: Sessions.Store): Path = path } def make(s: String): Context = if (s == ":") standard else dir(Path.explode(s)) } abstract class Context private { def enabled: Boolean def enabled(info: Sessions.Info): Boolean = enabled || info.browser_info def dir(store: Sessions.Store): Path = store.presentation_dir def dir(store: Sessions.Store, info: Sessions.Info): Path = dir(store) + Path.explode(info.chapter_session) } /** HTML presentation **/ /* maintain chapter index */ private val sessions_path = Path.basic(".sessions") private def read_sessions(dir: Path): List[(String, String)] = { val path = dir + sessions_path if (path.is_file) { import XML.Decode._ list(pair(string, string))(Symbol.decode_yxml(File.read(path))) } else Nil } private def write_sessions(dir: Path, sessions: List[(String, String)]): Unit = { import XML.Encode._ File.write(dir + sessions_path, YXML.string_of_body(list(pair(string, string))(sessions))) } def update_chapter_index( browser_info: Path, chapter: String, new_sessions: List[(String, String)]): Unit = { val dir = Isabelle_System.make_directory(browser_info + Path.basic(chapter)) val sessions0 = try { read_sessions(dir) } catch { case _: XML.Error => Nil } val sessions = (SortedMap.empty[String, String] ++ sessions0 ++ new_sessions).toList write_sessions(dir, sessions) val title = "Isabelle/" + chapter + " sessions" HTML.write_document(dir, "index.html", - List(HTML.title(title + " (" + Distribution.version + ")")), + List(HTML.title(title + Isabelle_System.isabelle_heading())), HTML.chapter(title) :: (if (sessions.isEmpty) Nil else List(HTML.div("sessions", List(HTML.description( sessions.map({ case (name, description) => val descr = Symbol.trim_blank_lines(description) (List(HTML.link(name + "/index.html", HTML.text(name))), if (descr == "") Nil else HTML.break ::: List(HTML.pre(HTML.text(descr)))) }))))))) } def make_global_index(browser_info: Path): Unit = { if (!(browser_info + Path.explode("index.html")).is_file) { Isabelle_System.make_directory(browser_info) Isabelle_System.copy_file(Path.explode("~~/lib/logo/isabelle.gif"), browser_info + Path.explode("isabelle.gif")) File.write(browser_info + Path.explode("index.html"), File.read(Path.explode("~~/lib/html/library_index_header.template")) + File.read(Path.explode("~~/lib/html/library_index_content.template")) + File.read(Path.explode("~~/lib/html/library_index_footer.template"))) } } /* present session */ val session_graph_path = Path.explode("session_graph.pdf") val readme_path = Path.explode("README.html") val files_path = Path.explode("files") def html_name(name: Document.Node.Name): String = name.theory_base_name + ".html" def theory_link(deps: Sessions.Deps, session0: String, name: Document.Node.Name, body: XML.Body): Option[XML.Tree] = { val session1 = deps(session0).theory_qualifier(name) val info0 = deps.sessions_structure.get(session0) val info1 = deps.sessions_structure.get(session1) if (info0.isDefined && info1.isDefined) { Some(HTML.link(info0.get.relative_path(info1.get) + html_name(name), body)) } else None } def session_html( resources: Resources, session: String, deps: Sessions.Deps, db_context: Sessions.Database_Context, progress: Progress = new Progress, verbose: Boolean = false, html_context: HTML_Context, elements: Elements, presentation: Context): Unit = { val info = deps.sessions_structure(session) val options = info.options val base = deps(session) val session_dir = presentation.dir(db_context.store, info) html_context.init_fonts(session_dir) Bytes.write(session_dir + session_graph_path, graphview.Graph_File.make_pdf(options, base.session_graph_display)) val documents = for { doc <- info.document_variants document <- db_context.input_database(session)(read_document(_, _, doc.name)) } yield { if (verbose) progress.echo("Presenting document " + session + "/" + doc.name) Bytes.write(session_dir + doc.path.pdf, document.pdf) doc } val view_links = { val deps_link = HTML.link(session_graph_path, HTML.text("theory dependencies")) val readme_links = if ((info.dir + readme_path).is_file) { Isabelle_System.copy_file(info.dir + readme_path, session_dir + readme_path) List(HTML.link(readme_path, HTML.text("README"))) } else Nil val document_links = documents.map(doc => HTML.link(doc.path.pdf, HTML.text(doc.name))) Library.separate(HTML.break ::: HTML.nl, (deps_link :: readme_links ::: document_links). map(link => HTML.text("View ") ::: List(link))).flatten } val theories: List[XML.Body] = { var seen_files = List.empty[(Path, String, Document.Node.Name)] def entity_link(props: Properties.T, body: XML.Body): Option[XML.Tree] = (props, props, props) match { case (Markup.Kind(Markup.THEORY), Markup.Name(theory), Position.Def_File(thy_file)) => val node_name = resources.file_node(Path.explode(thy_file), theory = theory) theory_link(deps, session, node_name, body) case _ => None } sealed case class Theory( name: Document.Node.Name, command: Command, files_html: List[(Path, XML.Tree)], html: XML.Tree) def read_theory(name: Document.Node.Name): Option[Theory] = { progress.expose_interrupt() if (verbose) progress.echo("Presenting theory " + name) for (command <- Build_Job.read_theory(db_context, resources, session, name.theory)) yield { val snapshot = Document.State.init.snippet(command) val files_html = for { (src_path, xml) <- snapshot.xml_markup_blobs(elements = elements.html) if xml.nonEmpty } yield { progress.expose_interrupt() if (verbose) progress.echo("Presenting file " + src_path) (src_path, html_context.source(make_html(elements, entity_link, xml))) } val html = html_context.source( make_html(elements, entity_link, snapshot.xml_markup(elements = elements.html))) Theory(name, command, files_html, html) } } for (thy <- Par_List.map(read_theory, base.session_theories).flatten) yield { val files = for { (src_path, file_html) <- thy.files_html } yield { val file_name = (files_path + src_path.squash.html).implode seen_files.find(p => p._1 == src_path || p._2 == file_name) match { case None => seen_files ::= (src_path, file_name, thy.name) case Some((_, _, thy_name1)) => error("Incoherent use of file name " + src_path + " as " + quote(file_name) + " in theory " + thy_name1 + " vs. " + thy.name) } val file_path = session_dir + Path.explode(file_name) html_context.init_fonts(file_path.dir) val file_title = "File " + Symbol.cartouche_decoded(src_path.implode_short) HTML.write_document(file_path.dir, file_path.file_name, List(HTML.title(file_title)), List(html_context.head(file_title), file_html)) List(HTML.link(file_name, HTML.text(file_title))) } val thy_title = "Theory " + thy.name.theory_base_name HTML.write_document(session_dir, html_name(thy.name), List(HTML.title(thy_title)), List(html_context.head(thy_title), thy.html)) List(HTML.link(html_name(thy.name), HTML.text(thy.name.theory_base_name) ::: (if (files.isEmpty) Nil else List(HTML.itemize(files))))) } } val title = "Session " + session HTML.write_document(session_dir, "index.html", - List(HTML.title(title + " (" + Distribution.version + ")")), + List(HTML.title(title + Isabelle_System.isabelle_heading())), html_context.head(title, List(HTML.par(view_links))) :: html_context.contents("Theories", theories)) } /** build documents **/ val session_tex_path = Path.explode("session.tex") def tex_name(name: Document.Node.Name): String = name.theory_base_name + ".tex" def document_tex_name(name: Document.Node.Name): String = "document/" + tex_name(name) class Build_Error(val log_lines: List[String], val message: String) extends Exn.User_Error(message) def build_documents( session: String, deps: Sessions.Deps, db_context: Sessions.Database_Context, progress: Progress = new Progress, output_sources: Option[Path] = None, output_pdf: Option[Path] = None, verbose: Boolean = false, verbose_latex: Boolean = false): List[Document_Output] = { /* session info */ val info = deps.sessions_structure(session) val hierarchy = deps.sessions_structure.hierarchy(session) val options = info.options val base = deps(session) val graph_pdf = graphview.Graph_File.make_pdf(options, base.session_graph_display) /* prepare document directory */ lazy val tex_files = for (name <- base.session_theories ::: base.document_theories) yield { val entry = db_context.get_export(hierarchy, name.theory, document_tex_name(name)) Path.basic(tex_name(name)) -> entry.uncompressed } def prepare_dir1(dir: Path, doc: Document_Variant): (Path, String) = { val doc_dir = dir + Path.basic(doc.name) Isabelle_System.make_directory(doc_dir) Isabelle_System.bash("isabelle latex -o sty", cwd = doc_dir.file).check File.write(doc_dir + Path.explode("isabelletags.sty"), doc.latex_sty) for ((base_dir, src) <- info.document_files) { Isabelle_System.copy_file_base(info.dir + base_dir, src, doc_dir) } File.write(doc_dir + session_tex_path, Library.terminate_lines( base.session_theories.map(name => "\\input{" + tex_name(name) + "}"))) for ((path, tex) <- tex_files) Bytes.write(doc_dir + path, tex) val root1 = "root_" + doc.name val root = if ((doc_dir + Path.explode(root1).tex).is_file) root1 else "root" (doc_dir, root) } def prepare_dir2(dir: Path, doc: Document_Variant): Unit = { val doc_dir = dir + Path.basic(doc.name) // non-deterministic, but irrelevant Bytes.write(doc_dir + session_graph_path, graph_pdf) } /* produce documents */ val documents = for (doc <- info.documents) yield { Isabelle_System.with_tmp_dir("document")(tmp_dir => { progress.echo("Preparing " + session + "/" + doc.name + " ...") val start = Time.now() // prepare sources val (doc_dir, root) = prepare_dir1(tmp_dir, doc) val digests = File.find_files(doc_dir.file, follow_links = true).map(SHA1.digest) val sources = SHA1.digest_set(digests) prepare_dir2(tmp_dir, doc) for (dir <- output_sources) { prepare_dir1(dir, doc) prepare_dir2(dir, doc) } // old document from database val old_document = for { old_doc <- db_context.input_database(session)(read_document(_, _, doc.name)) if old_doc.sources == sources } yield { Bytes.write(doc_dir + doc.path.pdf, old_doc.pdf) old_doc } old_document getOrElse { // bash scripts def root_bash(ext: String): String = Bash.string(root + "." + ext) def latex_bash(fmt: String = "pdf", ext: String = "tex"): String = "isabelle latex -o " + Bash.string(fmt) + " " + Bash.string(root + "." + ext) def bash(items: String*): Process_Result = progress.bash(items.mkString(" && "), cwd = doc_dir.file, echo = verbose_latex, watchdog = Time.seconds(0.5)) // prepare document val result = if ((doc_dir + Path.explode("build")).is_file) { bash("./build pdf " + Bash.string(doc.name)) } else { bash( latex_bash(), "{ [ ! -f " + root_bash("bib") + " ] || " + latex_bash("bbl") + "; }", "{ [ ! -f " + root_bash("idx") + " ] || " + latex_bash("idx") + "; }", latex_bash(), latex_bash()) } // result val root_pdf = Path.basic(root).pdf val result_path = doc_dir + root_pdf val log_lines = result.out_lines ::: result.err_lines val errors = (if (result.ok) Nil else List(result.err)) ::: Latex.latex_errors(doc_dir, root) ::: Bibtex.bibtex_errors(doc_dir, root) if (errors.nonEmpty) { val message = Exn.cat_message( (errors ::: List("Failed to build document " + quote(doc.name))):_*) throw new Build_Error(log_lines, message) } else if (!result_path.is_file) { val message = "Bad document result: expected to find " + root_pdf throw new Build_Error(log_lines, message) } else { val stop = Time.now() val timing = stop - start progress.echo("Finished " + session + "/" + doc.name + " (" + timing.message_hms + " elapsed time)") val log_xz = Bytes(cat_lines(log_lines)).compress() val pdf = Bytes.read(result_path) Document_Output(doc.name, sources, log_xz, pdf) } } }) } for (dir <- output_pdf; doc <- documents) { Isabelle_System.make_directory(dir) val path = dir + doc.path.pdf Bytes.write(path, doc.pdf) progress.echo("Document at " + path.absolute) } documents } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, args => { var output_sources: Option[Path] = None var output_pdf: Option[Path] = None var verbose_latex = false var dirs: List[Path] = Nil var options = Options.init() var verbose_build = false val getopts = Getopts( """ Usage: isabelle document [OPTIONS] SESSION Options are: -O DIR output directory for LaTeX sources and resulting PDF -P DIR output directory for resulting PDF -S DIR output directory for LaTeX sources -V verbose latex -d DIR include session directory -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose build Prepare the theory document of a session. """, "O:" -> (arg => { val dir = Path.explode(arg) output_sources = Some(dir) output_pdf = Some(dir) }), "P:" -> (arg => { output_pdf = Some(Path.explode(arg)) }), "S:" -> (arg => { output_sources = Some(Path.explode(arg)) }), "V" -> (_ => verbose_latex = true), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "o:" -> (arg => options = options + arg), "v" -> (_ => verbose_build = true)) val more_args = getopts(args) val session = more_args match { case List(a) => a case _ => getopts.usage() } val progress = new Console_Progress(verbose = verbose_build) val store = Sessions.store(options) progress.interrupt_handler { val res = Build.build(options, selection = Sessions.Selection.session(session), dirs = dirs, progress = progress, verbose = verbose_build) if (!res.ok) error("Failed to build session " + quote(session)) val deps = Sessions.load_structure(options + "document=pdf", dirs = dirs). selection_deps(Sessions.Selection.session(session)) if (output_sources.isEmpty && output_pdf.isEmpty) { progress.echo_warning("No output directory") } using(store.open_database_context())(db_context => build_documents(session, deps, db_context, progress = progress, output_sources = output_sources, output_pdf = output_pdf, verbose = true, verbose_latex = verbose_latex)) } }) } diff --git a/src/Pure/Tools/server.scala b/src/Pure/Tools/server.scala --- a/src/Pure/Tools/server.scala +++ b/src/Pure/Tools/server.scala @@ -1,598 +1,598 @@ /* Title: Pure/Tools/server.scala Author: Makarius Resident Isabelle servers. Message formats: - short message (single line): NAME ARGUMENT - long message (multiple lines): BYTE_LENGTH NAME ARGUMENT Argument formats: - Unit as empty string - XML.Elem in YXML notation - JSON.T in standard notation */ package isabelle import java.io.{BufferedInputStream, BufferedOutputStream, InputStreamReader, OutputStreamWriter, IOException} import java.net.{Socket, SocketException, SocketTimeoutException, ServerSocket, InetAddress} object Server { /* message argument */ object Argument { def is_name_char(c: Char): Boolean = Symbol.is_ascii_letter(c) || Symbol.is_ascii_digit(c) || c == '_' || c == '.' def split(msg: String): (String, String) = { val name = msg.takeWhile(is_name_char) val argument = msg.substring(name.length).dropWhile(Symbol.is_ascii_blank) (name, argument) } def print(arg: Any): String = arg match { case () => "" case t: XML.Elem => YXML.string_of_tree(t) case t: JSON.T => JSON.Format(t) } def parse(argument: String): Any = if (argument == "") () else if (YXML.detect_elem(argument)) YXML.parse_elem(argument) else JSON.parse(argument, strict = false) def unapply(argument: String): Option[Any] = try { Some(parse(argument)) } catch { case ERROR(_) => None } } /* input command */ type Command_Body = PartialFunction[(Context, Any), Any] abstract class Command(val command_name: String) { def command_body: Command_Body override def toString: String = command_name } class Commands(commands: Command*) extends Isabelle_System.Service { def entries: List[Command] = commands.toList } private lazy val command_table: Map[String, Command] = Isabelle_System.make_services(classOf[Commands]).flatMap(_.entries). foldLeft(Map.empty[String, Command]) { case (cmds, cmd) => val name = cmd.command_name if (cmds.isDefinedAt(name)) error("Duplicate Isabelle server command: " + quote(name)) else cmds + (name -> cmd) } /* output reply */ class Error(val message: String, val json: JSON.Object.T = JSON.Object.empty) extends RuntimeException(message) def json_error(exn: Throwable): JSON.Object.T = exn match { case e: Error => Reply.error_message(e.message) ++ e.json case ERROR(msg) => Reply.error_message(msg) case _ if Exn.is_interrupt(exn) => Reply.error_message(Exn.message(exn)) case _ => JSON.Object.empty } object Reply extends Enumeration { val OK, ERROR, FINISHED, FAILED, NOTE = Value def message(msg: String, kind: String = ""): JSON.Object.T = JSON.Object(Markup.KIND -> proper_string(kind).getOrElse(Markup.WRITELN), "message" -> msg) def error_message(msg: String): JSON.Object.T = message(msg, kind = Markup.ERROR) def unapply(msg: String): Option[(Reply.Value, Any)] = { if (msg == "") None else { val (name, argument) = Argument.split(msg) for { reply <- try { Some(withName(name)) } catch { case _: NoSuchElementException => None } arg <- Argument.unapply(argument) } yield (reply, arg) } } } /* handler: port, password, thread */ abstract class Handler(port0: Int) { val socket: ServerSocket = new ServerSocket(port0, 50, Server.localhost) def port: Int = socket.getLocalPort val password: String = UUID.random_string() override def toString: String = print(port, password) def handle(connection: Server.Connection): Unit private lazy val thread: Thread = Isabelle_Thread.fork(name = "server_handler") { var finished = false while (!finished) { Exn.capture(socket.accept) match { case Exn.Res(client) => Isabelle_Thread.fork(name = "client") { using(Connection(client))(connection => if (connection.read_password(password)) handle(connection)) } case Exn.Exn(_) => finished = true } } } def start(): Unit = thread def join: Unit = thread.join def stop(): Unit = { socket.close(); join } } /* socket connection */ object Connection { def apply(socket: Socket): Connection = new Connection(socket) } class Connection private(socket: Socket) extends AutoCloseable { override def toString: String = socket.toString def close(): Unit = socket.close() def set_timeout(t: Time): Unit = socket.setSoTimeout(t.ms.toInt) private val in = new BufferedInputStream(socket.getInputStream) private val out = new BufferedOutputStream(socket.getOutputStream) private val out_lock: AnyRef = new Object def tty_loop(): TTY_Loop = new TTY_Loop( new OutputStreamWriter(out), new InputStreamReader(in), writer_lock = out_lock) def read_password(password: String): Boolean = try { Byte_Message.read_line(in).map(_.text) == Some(password) } catch { case _: IOException => false } def read_message(): Option[String] = try { Byte_Message.read_line_message(in).map(_.text) } catch { case _: IOException => None } def await_close(): Unit = try { Byte_Message.read(in, 1); socket.close() } catch { case _: IOException => } def write_message(msg: String): Unit = out_lock.synchronized { Byte_Message.write_line_message(out, Bytes(UTF8.bytes(msg))) } def reply(r: Reply.Value, arg: Any): Unit = { val argument = Argument.print(arg) write_message(if (argument == "") r.toString else r.toString + " " + argument) } def reply_ok(arg: Any): Unit = reply(Reply.OK, arg) def reply_error(arg: Any): Unit = reply(Reply.ERROR, arg) def reply_error_message(message: String, more: JSON.Object.Entry*): Unit = reply_error(Reply.error_message(message) ++ more) def notify(arg: Any): Unit = reply(Reply.NOTE, arg) } /* context with output channels */ class Context private[Server](val server: Server, connection: Connection) extends AutoCloseable { context => def command_list: List[String] = command_table.keys.toList.sorted def reply(r: Reply.Value, arg: Any): Unit = connection.reply(r, arg) def notify(arg: Any): Unit = connection.notify(arg) def message(kind: String, msg: String, more: JSON.Object.Entry*): Unit = notify(Reply.message(msg, kind = kind) ++ more) def writeln(msg: String, more: JSON.Object.Entry*): Unit = message(Markup.WRITELN, msg, more:_*) def warning(msg: String, more: JSON.Object.Entry*): Unit = message(Markup.WARNING, msg, more:_*) def error_message(msg: String, more: JSON.Object.Entry*): Unit = message(Markup.ERROR, msg, more:_*) def progress(more: JSON.Object.Entry*): Connection_Progress = new Connection_Progress(context, more:_*) override def toString: String = connection.toString /* asynchronous tasks */ private val _tasks = Synchronized(Set.empty[Task]) def make_task(body: Task => JSON.Object.T): Task = { val task = new Task(context, body) _tasks.change(_ + task) task } def remove_task(task: Task): Unit = _tasks.change(_ - task) def cancel_task(id: UUID.T): Unit = _tasks.change(tasks => { tasks.find(task => task.id == id).foreach(_.cancel()); tasks }) def close(): Unit = { while(_tasks.change_result(tasks => { tasks.foreach(_.cancel()); (tasks.nonEmpty, tasks) })) { _tasks.value.foreach(_.join) } } } class Connection_Progress private[Server](context: Context, more: JSON.Object.Entry*) extends Progress { override def echo(msg: String): Unit = context.writeln(msg, more:_*) override def echo_warning(msg: String): Unit = context.warning(msg, more:_*) override def echo_error_message(msg: String): Unit = context.error_message(msg, more:_*) override def theory(theory: Progress.Theory): Unit = { val entries: List[JSON.Object.Entry] = List("theory" -> theory.theory, "session" -> theory.session) ::: (theory.percentage match { case None => Nil case Some(p) => List("percentage" -> p) }) context.writeln(theory.message, entries ::: more.toList:_*) } override def nodes_status(nodes_status: Document_Status.Nodes_Status): Unit = { val json = for ((name, node_status) <- nodes_status.present) yield name.json + ("status" -> node_status.json) context.notify(JSON.Object(Markup.KIND -> Markup.NODES_STATUS, Markup.NODES_STATUS -> json)) } override def toString: String = context.toString } class Task private[Server](val context: Context, body: Task => JSON.Object.T) { task => val id: UUID.T = UUID.random() val ident: JSON.Object.Entry = ("task" -> id.toString) val progress: Connection_Progress = context.progress(ident) def cancel(): Unit = progress.stop() private lazy val thread = Isabelle_Thread.fork(name = "server_task") { Exn.capture { body(task) } match { case Exn.Res(res) => context.reply(Reply.FINISHED, res + ident) case Exn.Exn(exn) => val err = json_error(exn) if (err.isEmpty) throw exn else context.reply(Reply.FAILED, err + ident) } progress.stop() context.remove_task(task) } def start: Unit = thread def join: Unit = thread.join } /* server info */ val localhost_name: String = "127.0.0.1" def localhost: InetAddress = InetAddress.getByName(localhost_name) def print_address(port: Int): String = localhost_name + ":" + port def print(port: Int, password: String): String = print_address(port) + " (password " + quote(password) + ")" object Info { private val Pattern = ("""server "([^"]*)" = \Q""" + localhost_name + """\E:(\d+) \(password "([^"]*)"\)""").r def parse(s: String): Option[Info] = s match { case Pattern(name, Value.Int(port), password) => Some(Info(name, port, password)) case _ => None } def apply(name: String, port: Int, password: String): Info = new Info(name, port, password) } class Info private(val name: String, val port: Int, val password: String) { def address: String = print_address(port) override def toString: String = "server " + quote(name) + " = " + print(port, password) def connection(): Connection = { val connection = Connection(new Socket(localhost, port)) connection.write_message(password) connection } def active: Boolean = try { using(connection())(connection => { connection.set_timeout(Time.seconds(2.0)) connection.read_message() match { case Some(Reply(Reply.OK, _)) => true case _ => false } }) } catch { case _: IOException => false case _: SocketException => false case _: SocketTimeoutException => false } } /* per-user servers */ val default_name = "isabelle" object Data { val database = Path.explode("$ISABELLE_HOME_USER/servers.db") val name = SQL.Column.string("name").make_primary_key val port = SQL.Column.int("port") val password = SQL.Column.string("password") val table = SQL.Table("isabelle_servers", List(name, port, password)) } def list(db: SQLite.Database): List[Info] = if (db.tables.contains(Data.table.name)) { db.using_statement(Data.table.select())(stmt => stmt.execute_query().iterator(res => Info( res.string(Data.name), res.int(Data.port), res.string(Data.password))).toList.sortBy(_.name)) } else Nil def find(db: SQLite.Database, name: String): Option[Info] = list(db).find(server_info => server_info.name == name && server_info.active) def init( name: String = default_name, port: Int = 0, existing_server: Boolean = false, log: Logger = No_Logger): (Info, Option[Server]) = { using(SQLite.open_database(Data.database))(db => { db.transaction { Isabelle_System.chmod("600", Data.database) db.create_table(Data.table) list(db).filterNot(_.active).foreach(server_info => db.using_statement(Data.table.delete(Data.name.where_equal(server_info.name)))( _.execute())) } db.transaction { find(db, name) match { case Some(server_info) => (server_info, None) case None => if (existing_server) error("Isabelle server " + quote(name) + " not running") val server = new Server(port, log) val server_info = Info(name, server.port, server.password) db.using_statement(Data.table.delete(Data.name.where_equal(name)))(_.execute()) db.using_statement(Data.table.insert())(stmt => { stmt.string(1) = server_info.name stmt.int(2) = server_info.port stmt.string(3) = server_info.password stmt.execute() }) server.start() (server_info, Some(server)) } } }) } def exit(name: String = default_name): Boolean = { using(SQLite.open_database(Data.database))(db => db.transaction { find(db, name) match { case Some(server_info) => using(server_info.connection())(_.write_message("shutdown")) while(server_info.active) { Time.seconds(0.05).sleep } true case None => false } }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("server", "manage resident Isabelle servers", Scala_Project.here, args => { var console = false var log_file: Option[Path] = None var operation_list = false var operation_exit = false var name = default_name var port = 0 var existing_server = false val getopts = Getopts(""" Usage: isabelle server [OPTIONS] Options are: -L FILE logging on FILE -c console interaction with specified server -l list servers (alternative operation) -n NAME explicit server name (default: """ + default_name + """) -p PORT explicit server port -s assume existing server, no implicit startup -x exit specified server (alternative operation) Manage resident Isabelle servers. """, "L:" -> (arg => log_file = Some(Path.explode(File.standard_path(arg)))), "c" -> (_ => console = true), "l" -> (_ => operation_list = true), "n:" -> (arg => name = arg), "p:" -> (arg => port = Value.Int.parse(arg)), "s" -> (_ => existing_server = true), "x" -> (_ => operation_exit = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() if (operation_list) { for { server_info <- using(SQLite.open_database(Data.database))(list) if server_info.active } Output.writeln(server_info.toString, stdout = true) } else if (operation_exit) { val ok = Server.exit(name) sys.exit(if (ok) 0 else 2) } else { val log = Logger.make(log_file) val (server_info, server) = init(name, port = port, existing_server = existing_server, log = log) Output.writeln(server_info.toString, stdout = true) if (console) { using(server_info.connection())(connection => connection.tty_loop().join) } server.foreach(_.join) } }) } class Server private(port0: Int, val log: Logger) extends Server.Handler(port0) { server => private val _sessions = Synchronized(Map.empty[UUID.T, Headless.Session]) def err_session(id: UUID.T): Nothing = error("No session " + Library.single_quote(id.toString)) def the_session(id: UUID.T): Headless.Session = _sessions.value.getOrElse(id, err_session(id)) def add_session(entry: (UUID.T, Headless.Session)): Unit = _sessions.change(_ + entry) def remove_session(id: UUID.T): Headless.Session = { _sessions.change_result(sessions => sessions.get(id) match { case Some(session) => (session, sessions - id) case None => err_session(id) }) } def shutdown(): Unit = { server.socket.close() val sessions = _sessions.change_result(sessions => (sessions, Map.empty)) for ((_, session) <- sessions) { try { val result = session.stop() if (!result.ok) log("Session shutdown failed: " + result.print_rc) } catch { case ERROR(msg) => log("Session shutdown failed: " + msg) } } } override def join: Unit = { super.join; shutdown() } override def handle(connection: Server.Connection): Unit = { using(new Server.Context(server, connection))(context => { connection.reply_ok( JSON.Object( "isabelle_id" -> Isabelle_System.isabelle_id(), - "isabelle_version" -> Distribution.version)) + "isabelle_name" -> Isabelle_System.isabelle_name())) var finished = false while (!finished) { connection.read_message() match { case None => finished = true case Some("") => context.notify("Command 'help' provides list of commands") case Some(msg) => val (name, argument) = Server.Argument.split(msg) Server.command_table.get(name) match { case Some(cmd) => argument match { case Server.Argument(arg) => if (cmd.command_body.isDefinedAt((context, arg))) { Exn.capture { cmd.command_body((context, arg)) } match { case Exn.Res(task: Server.Task) => connection.reply_ok(JSON.Object(task.ident)) task.start case Exn.Res(res) => connection.reply_ok(res) case Exn.Exn(exn) => val err = Server.json_error(exn) if (err.isEmpty) throw exn else connection.reply_error(err) } } else { connection.reply_error_message( "Bad argument for command " + Library.single_quote(name), "argument" -> argument) } case _ => connection.reply_error_message( "Malformed argument for command " + Library.single_quote(name), "argument" -> argument) } case None => connection.reply_error("Bad command " + Library.single_quote(name)) } } } }) } } diff --git a/src/Pure/build-jars b/src/Pure/build-jars --- a/src/Pure/build-jars +++ b/src/Pure/build-jars @@ -1,326 +1,325 @@ #!/usr/bin/env bash # # Author: Makarius # # build-jars - build Isabelle/Scala # # Requires proper Isabelle settings environment. ## sources declare -a SOURCES=( src/HOL/SPARK/Tools/spark.scala src/HOL/Tools/ATP/system_on_tptp.scala src/HOL/Tools/Nitpick/kodkod.scala src/Pure/Admin/afp.scala src/Pure/Admin/build_csdp.scala src/Pure/Admin/build_cygwin.scala src/Pure/Admin/build_doc.scala src/Pure/Admin/build_e.scala src/Pure/Admin/build_fonts.scala src/Pure/Admin/build_history.scala src/Pure/Admin/build_jcef.scala src/Pure/Admin/build_jdk.scala src/Pure/Admin/build_log.scala src/Pure/Admin/build_polyml.scala src/Pure/Admin/build_release.scala src/Pure/Admin/build_spass.scala src/Pure/Admin/build_sqlite.scala src/Pure/Admin/build_status.scala src/Pure/Admin/build_vampire.scala src/Pure/Admin/build_verit.scala src/Pure/Admin/build_zipperposition.scala src/Pure/Admin/check_sources.scala src/Pure/Admin/ci_profile.scala src/Pure/Admin/components.scala src/Pure/Admin/isabelle_cronjob.scala src/Pure/Admin/isabelle_devel.scala src/Pure/Admin/jenkins.scala src/Pure/Admin/other_isabelle.scala src/Pure/Concurrent/consumer_thread.scala src/Pure/Concurrent/counter.scala src/Pure/Concurrent/delay.scala src/Pure/Concurrent/event_timer.scala src/Pure/Concurrent/future.scala src/Pure/Concurrent/isabelle_thread.scala src/Pure/Concurrent/mailbox.scala src/Pure/Concurrent/par_list.scala src/Pure/Concurrent/synchronized.scala src/Pure/GUI/color_value.scala src/Pure/GUI/desktop_app.scala src/Pure/GUI/gui.scala src/Pure/GUI/gui_thread.scala src/Pure/GUI/popup.scala src/Pure/GUI/wrap_panel.scala src/Pure/General/antiquote.scala src/Pure/General/bytes.scala src/Pure/General/cache.scala src/Pure/General/codepoint.scala src/Pure/General/comment.scala src/Pure/General/completion.scala src/Pure/General/csv.scala src/Pure/General/date.scala src/Pure/General/exn.scala src/Pure/General/file.scala src/Pure/General/file_watcher.scala src/Pure/General/graph.scala src/Pure/General/graph_display.scala src/Pure/General/graphics_file.scala src/Pure/General/http.scala src/Pure/General/json.scala src/Pure/General/linear_set.scala src/Pure/General/logger.scala src/Pure/General/long_name.scala src/Pure/General/mailman.scala src/Pure/General/mercurial.scala src/Pure/General/multi_map.scala src/Pure/General/output.scala src/Pure/General/path.scala src/Pure/General/position.scala src/Pure/General/pretty.scala src/Pure/General/properties.scala src/Pure/General/rdf.scala src/Pure/General/scan.scala src/Pure/General/sha1.scala src/Pure/General/sql.scala src/Pure/General/ssh.scala src/Pure/General/symbol.scala src/Pure/General/time.scala src/Pure/General/timing.scala src/Pure/General/untyped.scala src/Pure/General/url.scala src/Pure/General/utf8.scala src/Pure/General/uuid.scala src/Pure/General/value.scala src/Pure/General/word.scala src/Pure/General/xz.scala src/Pure/Isar/document_structure.scala src/Pure/Isar/keyword.scala src/Pure/Isar/line_structure.scala src/Pure/Isar/outer_syntax.scala src/Pure/Isar/parse.scala src/Pure/Isar/token.scala src/Pure/ML/ml_console.scala src/Pure/ML/ml_lex.scala src/Pure/ML/ml_process.scala src/Pure/ML/ml_statistics.scala src/Pure/ML/ml_syntax.scala src/Pure/PIDE/byte_message.scala src/Pure/PIDE/command.scala src/Pure/PIDE/command_span.scala src/Pure/PIDE/document.scala src/Pure/PIDE/document_id.scala src/Pure/PIDE/document_status.scala src/Pure/PIDE/editor.scala src/Pure/PIDE/headless.scala src/Pure/PIDE/line.scala src/Pure/PIDE/markup.scala src/Pure/PIDE/markup_tree.scala src/Pure/PIDE/protocol.scala src/Pure/PIDE/protocol_handlers.scala src/Pure/PIDE/protocol_message.scala src/Pure/PIDE/prover.scala src/Pure/PIDE/query_operation.scala src/Pure/PIDE/rendering.scala src/Pure/PIDE/resources.scala src/Pure/PIDE/session.scala src/Pure/PIDE/text.scala src/Pure/PIDE/xml.scala src/Pure/PIDE/yxml.scala src/Pure/ROOT.scala src/Pure/System/bash.scala src/Pure/System/command_line.scala src/Pure/System/cygwin.scala - src/Pure/System/distribution.scala src/Pure/System/executable.scala src/Pure/System/getopts.scala src/Pure/System/isabelle_charset.scala src/Pure/System/isabelle_fonts.scala src/Pure/System/isabelle_platform.scala src/Pure/System/isabelle_process.scala src/Pure/System/isabelle_system.scala src/Pure/System/isabelle_tool.scala src/Pure/System/java_statistics.scala src/Pure/System/linux.scala src/Pure/System/mingw.scala src/Pure/System/numa.scala src/Pure/System/options.scala src/Pure/System/platform.scala src/Pure/System/posix_interrupt.scala src/Pure/System/process_result.scala src/Pure/System/progress.scala src/Pure/System/scala.scala src/Pure/System/system_channel.scala src/Pure/System/tty_loop.scala src/Pure/Thy/bibtex.scala src/Pure/Thy/export.scala src/Pure/Thy/export_theory.scala src/Pure/Thy/file_format.scala src/Pure/Thy/html.scala src/Pure/Thy/latex.scala src/Pure/Thy/presentation.scala src/Pure/Thy/sessions.scala src/Pure/Thy/thy_element.scala src/Pure/Thy/thy_header.scala src/Pure/Thy/thy_syntax.scala src/Pure/Tools/build.scala src/Pure/Tools/build_docker.scala src/Pure/Tools/build_job.scala src/Pure/Tools/check_keywords.scala src/Pure/Tools/debugger.scala src/Pure/Tools/doc.scala src/Pure/Tools/dump.scala src/Pure/Tools/fontforge.scala src/Pure/Tools/java_monitor.scala src/Pure/Tools/logo.scala src/Pure/Tools/main.scala src/Pure/Tools/mkroot.scala src/Pure/Tools/phabricator.scala src/Pure/Tools/print_operation.scala src/Pure/Tools/profiling_report.scala src/Pure/Tools/scala_project.scala src/Pure/Tools/server.scala src/Pure/Tools/server_commands.scala src/Pure/Tools/simplifier_trace.scala src/Pure/Tools/spell_checker.scala src/Pure/Tools/task_statistics.scala src/Pure/Tools/update.scala src/Pure/Tools/update_cartouches.scala src/Pure/Tools/update_comments.scala src/Pure/Tools/update_header.scala src/Pure/Tools/update_then.scala src/Pure/Tools/update_theorems.scala src/Pure/library.scala src/Pure/pure_thy.scala src/Pure/term.scala src/Pure/term_xml.scala src/Pure/thm_name.scala src/Tools/Graphview/graph_file.scala src/Tools/Graphview/graph_panel.scala src/Tools/Graphview/graphview.scala src/Tools/Graphview/layout.scala src/Tools/Graphview/main_panel.scala src/Tools/Graphview/metrics.scala src/Tools/Graphview/model.scala src/Tools/Graphview/mutator.scala src/Tools/Graphview/mutator_dialog.scala src/Tools/Graphview/mutator_event.scala src/Tools/Graphview/popups.scala src/Tools/Graphview/shapes.scala src/Tools/Graphview/tree_panel.scala src/Tools/VSCode/src/build_vscode.scala src/Tools/VSCode/src/channel.scala src/Tools/VSCode/src/dynamic_output.scala src/Tools/VSCode/src/language_server.scala src/Tools/VSCode/src/lsp.scala src/Tools/VSCode/src/preview_panel.scala src/Tools/VSCode/src/state_panel.scala src/Tools/VSCode/src/textmate_grammar.scala src/Tools/VSCode/src/vscode_model.scala src/Tools/VSCode/src/vscode_rendering.scala src/Tools/VSCode/src/vscode_resources.scala src/Tools/VSCode/src/vscode_spell_checker.scala ) ## diagnostics PRG="$(basename "$0")" function usage() { echo echo "Usage: isabelle $PRG [OPTIONS]" echo echo " Options are:" echo " -f fresh build" echo exit 1 } function fail() { echo "$1" >&2 exit 2 } [ -z "$ISABELLE_HOME" ] && fail "Missing Isabelle settings environment" ## process command line # options FRESH="" while getopts "f" OPT do case "$OPT" in f) FRESH=true ;; \?) usage ;; esac done shift $(($OPTIND - 1)) # args [ "$#" -ne 0 ] && usage ## target TARGET_DIR="lib/classes" TARGET_JAR="$TARGET_DIR/Pure.jar" TARGET_SHASUM="$TARGET_DIR/Pure.shasum" function target_shasum() { shasum -a1 -b "$TARGET_JAR" "${SOURCES[@]}" 2>/dev/null } function target_clean() { rm -rf "$TARGET_DIR" } [ -n "$FRESH" ] && target_clean ## build target_shasum | cmp "$TARGET_SHASUM" >/dev/null 2>/dev/null if [ "$?" -ne 0 ]; then echo "### Building Isabelle/Scala ..." target_clean BUILD_DIR="$TARGET_DIR/build" mkdir -p "$BUILD_DIR" ( export CLASSPATH="$(platform_path "$ISABELLE_CLASSPATH")" isabelle_scala scalac $ISABELLE_SCALAC_OPTIONS \ -d "$BUILD_DIR" "${SOURCES[@]}" ) || fail "Failed to compile sources" CHARSET_SERVICE="META-INF/services/java.nio.charset.spi.CharsetProvider" mkdir -p "$BUILD_DIR/$(dirname "$CHARSET_SERVICE")" echo isabelle.Isabelle_Charset_Provider > "$BUILD_DIR/$CHARSET_SERVICE" cp "$ISABELLE_HOME/lib/logo/isabelle_transparent-32.gif" "$BUILD_DIR/isabelle/." cp "$ISABELLE_HOME/lib/logo/isabelle_transparent.gif" "$BUILD_DIR/isabelle/." isabelle_jdk jar -c -f "$(platform_path "$TARGET_JAR")" -e isabelle.Main \ -C "$BUILD_DIR" META-INF \ -C "$BUILD_DIR" isabelle || fail "Failed to produce $TARGET_JAR" rm -rf "$BUILD_DIR" target_shasum > "$TARGET_SHASUM" fi diff --git a/src/Tools/VSCode/src/language_server.scala b/src/Tools/VSCode/src/language_server.scala --- a/src/Tools/VSCode/src/language_server.scala +++ b/src/Tools/VSCode/src/language_server.scala @@ -1,562 +1,562 @@ /* Title: Tools/VSCode/src/language_server.scala Author: Makarius Server for VS Code Language Server Protocol 2.0/3.0, see also https://github.com/Microsoft/language-server-protocol https://github.com/Microsoft/language-server-protocol/blob/master/protocol.md PIDE protocol extensions depend on system option "vscode_pide_extensions". */ package isabelle.vscode import isabelle._ import java.io.{PrintStream, OutputStream, File => JFile} import scala.annotation.tailrec import scala.collection.mutable object Language_Server { type Editor = isabelle.Editor[Unit] /* Isabelle tool wrapper */ private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC") val isabelle_tool = Isabelle_Tool("vscode_server", "VSCode Language Server for PIDE", Scala_Project.here, args => { try { var logic_ancestor: Option[String] = None var log_file: Option[Path] = None var logic_requirements = false var dirs: List[Path] = Nil var include_sessions: List[String] = Nil var logic = default_logic var modes: List[String] = Nil var options = Options.init() var verbose = false val getopts = Getopts(""" Usage: isabelle vscode_server [OPTIONS] Options are: -A NAME ancestor session for option -R (default: parent) -L FILE logging on FILE -R NAME build image with requirements from other sessions -d DIR include session directory -i NAME include session in name-space of theories -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(default_logic) + """) -m MODE add print mode for output -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose logging Run the VSCode Language Server protocol (JSON RPC) over stdin/stdout. """, "A:" -> (arg => logic_ancestor = Some(arg)), "L:" -> (arg => log_file = Some(Path.explode(File.standard_path(arg)))), "R:" -> (arg => { logic = arg; logic_requirements = true }), "d:" -> (arg => dirs = dirs ::: List(Path.explode(File.standard_path(arg)))), "i:" -> (arg => include_sessions = include_sessions ::: List(arg)), "l:" -> (arg => logic = arg), "m:" -> (arg => modes = arg :: modes), "o:" -> (arg => options = options + arg), "v" -> (_ => verbose = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val log = Logger.make(log_file) val channel = new Channel(System.in, System.out, log, verbose) val server = new Language_Server(channel, options, session_name = logic, session_dirs = dirs, include_sessions = include_sessions, session_ancestor = logic_ancestor, session_requirements = logic_requirements, modes = modes, log = log) // prevent spurious garbage on the main protocol channel val orig_out = System.out try { System.setOut(new PrintStream(new OutputStream { def write(n: Int): Unit = {} })) server.start() } finally { System.setOut(orig_out) } } catch { case exn: Throwable => val channel = new Channel(System.in, System.out, No_Logger) channel.error_message(Exn.message(exn)) throw(exn) } }) } class Language_Server( val channel: Channel, options: Options, session_name: String = Language_Server.default_logic, session_dirs: List[Path] = Nil, include_sessions: List[String] = Nil, session_ancestor: Option[String] = None, session_requirements: Boolean = false, modes: List[String] = Nil, log: Logger = No_Logger) { server => /* prover session */ private val session_ = Synchronized(None: Option[Session]) def session: Session = session_.value getOrElse error("Server inactive") def resources: VSCode_Resources = session.resources.asInstanceOf[VSCode_Resources] def rendering_offset(node_pos: Line.Node_Position): Option[(VSCode_Rendering, Text.Offset)] = for { model <- resources.get_model(new JFile(node_pos.name)) rendering = model.rendering() offset <- model.content.doc.offset(node_pos.pos) } yield (rendering, offset) private val dynamic_output = Dynamic_Output(server) /* input from client or file-system */ private val delay_input: Delay = Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger) { resources.flush_input(session, channel) } private val delay_load: Delay = Delay.last(options.seconds("vscode_load_delay"), channel.Error_Logger) { val (invoke_input, invoke_load) = resources.resolve_dependencies(session, editor, file_watcher) if (invoke_input) delay_input.invoke() if (invoke_load) delay_load.invoke() } private val file_watcher = File_Watcher(sync_documents, options.seconds("vscode_load_delay")) private def close_document(file: JFile): Unit = { if (resources.close_model(file)) { file_watcher.register_parent(file) sync_documents(Set(file)) delay_input.invoke() delay_output.invoke() } } private def sync_documents(changed: Set[JFile]): Unit = { resources.sync_models(changed) delay_input.invoke() delay_output.invoke() } private def change_document( file: JFile, version: Long, changes: List[LSP.TextDocumentChange]): Unit = { val norm_changes = new mutable.ListBuffer[LSP.TextDocumentChange] @tailrec def norm(chs: List[LSP.TextDocumentChange]): Unit = { if (chs.nonEmpty) { val (full_texts, rest1) = chs.span(_.range.isEmpty) val (edits, rest2) = rest1.span(_.range.nonEmpty) norm_changes ++= full_texts norm_changes ++= edits.sortBy(_.range.get.start)(Line.Position.Ordering).reverse norm(rest2) } } norm(changes) norm_changes.foreach(change => resources.change_model(session, editor, file, version, change.text, change.range)) delay_input.invoke() delay_output.invoke() } /* caret handling */ private val delay_caret_update: Delay = Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger) { session.caret_focus.post(Session.Caret_Focus) } private def update_caret(caret: Option[(JFile, Line.Position)]): Unit = { resources.update_caret(caret) delay_caret_update.invoke() delay_input.invoke() } /* preview */ private lazy val preview_panel = new Preview_Panel(resources) private lazy val delay_preview: Delay = Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger) { if (preview_panel.flush(channel)) delay_preview.invoke() } private def request_preview(file: JFile, column: Int): Unit = { preview_panel.request(file, column) delay_preview.invoke() } /* output to client */ private val delay_output: Delay = Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger) { if (resources.flush_output(channel)) delay_output.invoke() } def update_output(changed_nodes: Iterable[JFile]): Unit = { resources.update_output(changed_nodes) delay_output.invoke() } def update_output_visible(): Unit = { resources.update_output_visible() delay_output.invoke() } private val prover_output = Session.Consumer[Session.Commands_Changed](getClass.getName) { case changed => update_output(changed.nodes.toList.map(resources.node_file(_))) } private val syslog_messages = Session.Consumer[Prover.Output](getClass.getName) { case output => channel.log_writeln(resources.output_xml(output.message)) } /* init and exit */ def init(id: LSP.Id): Unit = { def reply_ok(msg: String): Unit = { channel.write(LSP.Initialize.reply(id, "")) channel.writeln(msg) } def reply_error(msg: String): Unit = { channel.write(LSP.Initialize.reply(id, msg)) channel.error_message(msg) } val try_session = try { val base_info = Sessions.base_info( options, session_name, dirs = session_dirs, include_sessions = include_sessions, session_ancestor = session_ancestor, session_requirements = session_requirements).check def build(no_build: Boolean = false): Build.Results = Build.build(options, selection = Sessions.Selection.session(base_info.session), build_heap = true, no_build = no_build, dirs = session_dirs, infos = base_info.infos) if (!build(no_build = true).ok) { val start_msg = "Build started for Isabelle/" + base_info.session + " ..." val fail_msg = "Session build failed -- prover process remains inactive!" val progress = channel.progress(verbose = true) progress.echo(start_msg); channel.writeln(start_msg) if (!build().ok) { progress.echo(fail_msg); error(fail_msg) } } val resources = new VSCode_Resources(options, base_info, log) { override def commit(change: Session.Change): Unit = if (change.deps_changed || undefined_blobs(change.version.nodes).nonEmpty) delay_load.invoke() } val session_options = options.bool("editor_output_state") = true val session = new Session(session_options, resources) Some((base_info, session)) } catch { case ERROR(msg) => reply_error(msg); None } for ((base_info, session) <- try_session) { session_.change(_ => Some(session)) session.commands_changed += prover_output session.syslog_messages += syslog_messages dynamic_output.init() try { Isabelle_Process(session, options, base_info.sessions_structure, Sessions.store(options), modes = modes, logic = base_info.session).await_startup() - reply_ok("Welcome to Isabelle/" + base_info.session + " (" + Distribution.version + ")") + reply_ok("Welcome to Isabelle/" + base_info.session + Isabelle_System.isabelle_heading()) } catch { case ERROR(msg) => reply_error(msg) } } } def shutdown(id: LSP.Id): Unit = { def reply(err: String): Unit = channel.write(LSP.Shutdown.reply(id, err)) session_.change({ case Some(session) => session.commands_changed -= prover_output session.syslog_messages -= syslog_messages dynamic_output.exit() delay_load.revoke() file_watcher.shutdown() delay_input.revoke() delay_output.revoke() delay_caret_update.revoke() delay_preview.revoke() val result = session.stop() if (result.ok) reply("") else reply("Prover shutdown failed: " + result.rc) None case None => reply("Prover inactive") None }) } def exit(): Unit = { log("\n") sys.exit(if (session_.value.isDefined) 2 else 0) } /* completion */ def completion(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = (for ((rendering, offset) <- rendering_offset(node_pos)) yield rendering.completion(node_pos.pos, offset)) getOrElse Nil channel.write(LSP.Completion.reply(id, result)) } /* spell-checker dictionary */ def update_dictionary(include: Boolean, permanent: Boolean): Unit = { for { spell_checker <- resources.spell_checker.get caret <- resources.get_caret() rendering = caret.model.rendering() range = rendering.before_caret_range(caret.offset) Text.Info(_, word) <- Spell_Checker.current_word(rendering, range) } { spell_checker.update(word, include, permanent) update_output_visible() } } def reset_dictionary(): Unit = { for (spell_checker <- resources.spell_checker.get) { spell_checker.reset() update_output_visible() } } /* hover */ def hover(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = for { (rendering, offset) <- rendering_offset(node_pos) info <- rendering.tooltips(VSCode_Rendering.tooltip_elements, Text.Range(offset, offset + 1)) } yield { val range = rendering.model.content.doc.range(info.range) val contents = info.info.map(t => LSP.MarkedString(resources.output_pretty_tooltip(List(t)))) (range, contents) } channel.write(LSP.Hover.reply(id, result)) } /* goto definition */ def goto_definition(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = (for ((rendering, offset) <- rendering_offset(node_pos)) yield rendering.hyperlinks(Text.Range(offset, offset + 1))) getOrElse Nil channel.write(LSP.GotoDefinition.reply(id, result)) } /* document highlights */ def document_highlights(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = (for ((rendering, offset) <- rendering_offset(node_pos)) yield { val model = rendering.model rendering.caret_focus_ranges(Text.Range(offset, offset + 1), model.content.text_range) .map(r => LSP.DocumentHighlight.text(model.content.doc.range(r))) }) getOrElse Nil channel.write(LSP.DocumentHighlights.reply(id, result)) } /* main loop */ def start(): Unit = { log("Server started " + Date.now()) def handle(json: JSON.T): Unit = { try { json match { case LSP.Initialize(id) => init(id) case LSP.Initialized(()) => case LSP.Shutdown(id) => shutdown(id) case LSP.Exit(()) => exit() case LSP.DidOpenTextDocument(file, _, version, text) => change_document(file, version, List(LSP.TextDocumentChange(None, text))) delay_load.invoke() case LSP.DidChangeTextDocument(file, version, changes) => change_document(file, version, changes) case LSP.DidCloseTextDocument(file) => close_document(file) case LSP.Completion(id, node_pos) => completion(id, node_pos) case LSP.Include_Word(()) => update_dictionary(true, false) case LSP.Include_Word_Permanently(()) => update_dictionary(true, true) case LSP.Exclude_Word(()) => update_dictionary(false, false) case LSP.Exclude_Word_Permanently(()) => update_dictionary(false, true) case LSP.Reset_Words(()) => reset_dictionary() case LSP.Hover(id, node_pos) => hover(id, node_pos) case LSP.GotoDefinition(id, node_pos) => goto_definition(id, node_pos) case LSP.DocumentHighlights(id, node_pos) => document_highlights(id, node_pos) case LSP.Caret_Update(caret) => update_caret(caret) case LSP.State_Init(()) => State_Panel.init(server) case LSP.State_Exit(id) => State_Panel.exit(id) case LSP.State_Locate(id) => State_Panel.locate(id) case LSP.State_Update(id) => State_Panel.update(id) case LSP.State_Auto_Update(id, enabled) => State_Panel.auto_update(id, enabled) case LSP.Preview_Request(file, column) => request_preview(file, column) case LSP.Symbols_Request(()) => channel.write(LSP.Symbols()) case _ => if (!LSP.ResponseMessage.is_empty(json)) log("### IGNORED") } } catch { case exn: Throwable => channel.log_error_message(Exn.message(exn)) } } @tailrec def loop(): Unit = { channel.read() match { case Some(json) => json match { case bulk: List[_] => bulk.foreach(handle) case _ => handle(json) } loop() case None => log("### TERMINATE") } } loop() } /* abstract editor operations */ object editor extends Language_Server.Editor { /* session */ override def session: Session = server.session override def flush(): Unit = resources.flush_input(session, channel) override def invoke(): Unit = delay_input.invoke() /* current situation */ override def current_node(context: Unit): Option[Document.Node.Name] = resources.get_caret().map(_.model.node_name) override def current_node_snapshot(context: Unit): Option[Document.Snapshot] = resources.get_caret().map(_.model.snapshot()) override def node_snapshot(name: Document.Node.Name): Document.Snapshot = { resources.get_model(name) match { case Some(model) => model.snapshot() case None => session.snapshot(name) } } def current_command(snapshot: Document.Snapshot): Option[Command] = { resources.get_caret() match { case Some(caret) => snapshot.current_command(caret.node_name, caret.offset) case None => None } } override def current_command(context: Unit, snapshot: Document.Snapshot): Option[Command] = current_command(snapshot) /* overlays */ override def node_overlays(name: Document.Node.Name): Document.Node.Overlays = resources.node_overlays(name) override def insert_overlay(command: Command, fn: String, args: List[String]): Unit = resources.insert_overlay(command, fn, args) override def remove_overlay(command: Command, fn: String, args: List[String]): Unit = resources.remove_overlay(command, fn, args) /* hyperlinks */ override def hyperlink_command( focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic, offset: Symbol.Offset = 0): Option[Hyperlink] = { if (snapshot.is_outdated) None else snapshot.find_command_position(id, offset).map(node_pos => new Hyperlink { def follow(unit: Unit): Unit = channel.write(LSP.Caret_Update(node_pos, focus)) }) } /* dispatcher thread */ override def assert_dispatcher[A](body: => A): A = session.assert_dispatcher(body) override def require_dispatcher[A](body: => A): A = session.require_dispatcher(body) override def send_dispatcher(body: => Unit): Unit = session.send_dispatcher(body) override def send_wait_dispatcher(body: => Unit): Unit = session.send_wait_dispatcher(body) } }