upgrade to Isabelle2017 and afp-2017

git-svn-id: https://projects.brucker.ch/hol-testgen/svn/HOL-TestGen/trunk/hol-testgen@13265 3260e6d1-4efc-4170-b0a7-36055960796d
This commit is contained in:
Frédéric Tuong 2018-02-05 14:26:56 +00:00
parent 50c90536d4
commit 208bb30780
27 changed files with 293 additions and 446 deletions

View File

@ -54,7 +54,7 @@ Isabelle (version 3.6) mentions:
of their own. No starting spaces, nothing after it."
In particular, it is not advised to put these tags in a single line:
\isatagafp ... \endisatagafp % wrong
otherwise as side effects some parts occuring after these tags may be
otherwise as side effects some parts occurring after these tags may be
skipped. The recommanded solution is to always write each tag in a
separate line:
\isatagafp
@ -78,6 +78,7 @@ are close to not debug-able.
List of Isabelle versions to use depending on revisions:
=========================================================
2018/02/05 revision 13265: Isabelle2017 (October 2017)
2018/01/29 revision 13259: Isabelle2016-1 (December 2016)
2016/02/22 revision 12439: Isabelle2016 (February 2016)
2015/06/11 revision 11691: Isabelle2015 (May 2015)

6
ROOT
View File

@ -101,7 +101,7 @@ session "OCL" in "src" = HOL +
(******************************************************)
session "OCL-all-dirty" in "src" = HOL +
session "OCL-all-dirty" in "src" = "HOL-Library" +
description {* Featherweight OCL (Long and Dirty) *}
options [quick_and_dirty,document=pdf,document_output=document_generated,
document_variants="document=afp,-annexa,-noexample",
@ -164,7 +164,7 @@ session "OCL-all-dirty" in "src" = HOL +
(******************************************************)
session "FOCL" in "src" = HOL +
session "FOCL" in "src" = "HOL-Library" +
description {* Featherweight OCL (Compiler) *}
options [document=pdf,document_output=document_generated,
document_variants="document=noexample,-afp,-annexa",
@ -189,7 +189,7 @@ session "FOCL" in "src" = HOL +
"root.tex"
"FOCL_Syntax.tex"
session "FOCL-dirty" in "src" = HOL +
session "FOCL-dirty" in "src" = "HOL-Library" +
description {* Featherweight OCL (Compiler) *}
options [quick_and_dirty,document=pdf,document_output=document_generated,
document_variants="document=noexample,-afp,-annexa",

View File

@ -637,7 +637,7 @@ by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny
section{* OclAllInstances *}
text{* To denote OCL-types occuring in OCL expressions syntactically---as, for example, as
text{* To denote OCL-types occurring in OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.'' *}

View File

@ -626,7 +626,7 @@ by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny
section{* OclAllInstances *}
text{* To denote OCL-types occuring in OCL expressions syntactically---as, for example, as
text{* To denote OCL-types occurring in OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.'' *}

View File

@ -122,7 +122,7 @@ like validity, not emptiness, finiteness...
Since the first hypothesis of @{text comp_fun_commute.fold_insert} is too general,
in order to replace it by another weaker locale we have the choice between
reusing the @{term comp_fun_commute} locale or whether completely defining a new locale.
Because elements occuring in the type of @{term Finite_Set.fold_graph} are represented in polymorphic form,
Because elements occurring in the type of @{term Finite_Set.fold_graph} are represented in polymorphic form,
the folding on a value-proposition couple would be possible in a type system with dependent types.
But without the dependent typing facility, we choose to give the well-defined properties
to each functions in a duplicated version of @{term comp_fun_commute}. *}

View File

@ -409,7 +409,7 @@ lemma mbind_unit [simp]:
text{* The characteristic property of @{term mbind} --- which distinguishes it from
@{text mbind'} defined in the sequel --- is that it never fails; it ``swallows'' internal
errors occuring during the computation. *}
errors occurring during the computation. *}
lemma mbind_nofailure [simp]:
"mbind S f \<sigma> \<noteq> None"
apply(rule_tac x=\<sigma> in spec)

View File

@ -47,7 +47,7 @@
chapter{* Formalization I: OCL Types and Core Definitions \label{sec:focl-types}*}
theory UML_Types
imports Transcendental (* Testing *)
imports HOL.Transcendental (* Testing *)
keywords "Assert" :: thy_decl
and "Assert_local" :: thy_decl
begin
@ -624,9 +624,9 @@ fun outer_syntax_command command_spec theory in_local =
Outer_Syntax.command command_spec "assert that the given specification is true"
(Parse.term >> (fn elems_concl => theory (fn thy =>
case
lemma "code_unfold" (Specification.theorem true)
lemma "nbe" (Specification.theorem true)
(fn lthy =>
let val expr = Value_Command.value lthy (Syntax.read_term lthy elems_concl)
let val expr = Nbe.dynamic_value lthy (Syntax.read_term lthy elems_concl)
val thy = Proof_Context.theory_of lthy
open HOLogic in
if Sign.typ_equiv thy (fastype_of expr, @{typ "prop"}) then

View File

@ -71,7 +71,7 @@ text{* Our notion of typed bag goes beyond the usual notion of a finite executab
is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means
we can have in Featherweight OCL Bags containing all possible elements of a type, not only
those (finite) ones representable in a state. This holds for base types as well as class types,
although the notion for class-types --- involving object id's not occuring in a state ---
although the notion for class-types --- involving object id's not occurring in a state ---
requires some care.
In a world with @{term invalid} and @{term null}, there are two notions extensions possible:

View File

@ -63,7 +63,7 @@ text{* Our notion of typed set goes beyond the usual notion of a finite executab
is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means
we can have in Featherweight OCL Sets containing all possible elements of a type, not only
those (finite) ones representable in a state. This holds for base types as well as class types,
although the notion for class-types --- involving object id's not occuring in a state ---
although the notion for class-types --- involving object id's not occurring in a state ---
requires some care.
In a world with @{term invalid} and @{term null}, there are two notions extensions possible:

View File

@ -157,7 +157,7 @@ lazy_text\<open>For each Class \emph{C}, we will have a casting operation \inlin
lazy_text\<open>Thus, since we have two class-types in our concrete class hierarchy, we have
two operations to declare and to provide two overloading definitions for the two static types.
\<close>
lazy_text\<open>To denote OCL-types occuring in OCL expressions syntactically---as, for example, as
lazy_text\<open>To denote OCL-types occurring in OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.''\<close>
lazy_text\<open>\label{sec:edm-accessors}\<close>
@ -237,7 +237,7 @@ lazy_text\<open>For each Class \emph{C}, we will have a casting operation \inlin
lazy_text\<open>Thus, since we have two class-types in our concrete class hierarchy, we have
two operations to declare and to provide two overloading definitions for the two static types.
\<close>
lazy_text\<open>To denote OCL-types occuring in OCL expressions syntactically---as, for example, as
lazy_text\<open>To denote OCL-types occurring in OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.''\<close>
lazy_text\<open>\label{sec:eam-accessors}\<close>

View File

@ -434,7 +434,7 @@ two operations to declare and to provide two overloading definitions for the two
, [ section \<open>OclAllInstances\<close>
, txt'' [ \<open>
To denote \OCL-types occuring in \OCL expressions syntactically---as, for example, as
To denote \OCL-types occurring in \OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.'' \<close> ]
, PRINT_allinst_def_id

View File

@ -200,7 +200,7 @@ fun check l_oid l =
META.check_export_code
(writeln o Mi)
(warning o Mi)
(writeln o Markup.markup Markup.bad o Mi)
(fn s => writeln (Markup.markup (Markup.bad ()) (Mi s)))
(error o To_string0)
(Ml (Mp I Me) l_oid)
((META.SS_base o META.ST) l)
@ -1042,7 +1042,7 @@ fun export_code_tmp_file seris g =
else
Isabelle_System.with_tmp_file tmp_name (Deep0.Find.ext ml_compiler))
(fn filename =>
g (((((ml_compiler, ml_module), Path.implode filename), export_arg) :: accu)))
g (((((ml_compiler, ml_module), (Path.implode filename, Position.none)), export_arg) :: accu)))
end))
seris
(fn f => f [])
@ -1063,7 +1063,7 @@ fun export_code_cmd' seris tmp_export_code f_err raw_cs thy =
let val v = Deep0.apply_hs_code_identifiers Deep0.Export_code_env.Haskell.argument thy in
if mem_scala then Code_printing.apply_code_printing v else v end) in
List_mapi
(fn i => fn seri => case seri of (((ml_compiler, _), filename), _) =>
(fn i => fn seri => case seri of (((ml_compiler, _), (filename, _)), _) =>
let val (l, (out, err)) =
Deep0.Find.build
ml_compiler
@ -1240,12 +1240,13 @@ fun f_command l_mode =
fun mk_fic s = Path.append tmp_export_code (Path.make [s])
val () = Deep0.Find.check_compil ml_compiler ()
val () = Isabelle_System.mkdirs tmp_export_code in
((( (ml_compiler, ml_module)
, Path.implode (if Deep0.Find.export_mode ml_compiler = Deep0.Export_code_env.Directory then
tmp_export_code
else
mk_fic (Deep0.Find.function ml_compiler (Deep0.Find.ext ml_compiler))))
, export_arg), mk_fic)
(( ( (ml_compiler, ml_module)
, ( Path.implode (if Deep0.Find.export_mode ml_compiler = Deep0.Export_code_env.Directory then
tmp_export_code
else
mk_fic (Deep0.Find.function ml_compiler (Deep0.Find.ext ml_compiler)))
, Position.none))
, export_arg), mk_fic)
end)
(List.filter (fn (("self", _), _) => false | _ => true) (#seri_args i_deep))
val _ =
@ -1318,7 +1319,7 @@ let open Generation_mode
:: []))
|> Deep.export_code_cmd' seri_args
tmp_export_code
(fn (((_, _), msg), _) => fn err => if err <> 0 then error msg else ())
(fn (((_, _), (msg, _)), _) => fn err => if err <> 0 then error msg else ())
[name_main]
end
in
@ -1363,7 +1364,7 @@ local
fun fold_thy_shallow f =
META.fold_thy_shallow
(fn f => f () handle ERROR e =>
( warning "Shallow Backtracking: (true) Isabelle declarations occuring among the META-simulated ones are ignored (if any)"
( warning "Shallow Backtracking: (true) Isabelle declarations occurring among the META-simulated ones are ignored (if any)"
(* TODO automatically determine if there is such Isabelle declarations,
for raising earlier a specific error message *)
; error e))

View File

@ -224,7 +224,7 @@ fun check l_oid l =
META.check_export_code
(writeln o Mi)
(warning o Mi)
(writeln o Markup.markup Markup.bad o Mi)
(fn s => writeln (Markup.markup (Markup.bad ()) (Mi s)))
(error o To_string0)
(Ml (Mp I Me) l_oid)
((META.SS_base o META.ST) l)
@ -1066,7 +1066,7 @@ fun export_code_tmp_file seris g =
else
Isabelle_System.with_tmp_file tmp_name (Deep0.Find.ext ml_compiler))
(fn filename =>
g (((((ml_compiler, ml_module), Path.implode filename), export_arg) :: accu)))
g (((((ml_compiler, ml_module), (Path.implode filename, Position.none)), export_arg) :: accu)))
end))
seris
(fn f => f [])
@ -1087,7 +1087,7 @@ fun export_code_cmd' seris tmp_export_code f_err raw_cs thy =
let val v = Deep0.apply_hs_code_identifiers Deep0.Export_code_env.Haskell.argument thy in
if mem_scala then Code_printing.apply_code_printing v else v end) in
List_mapi
(fn i => fn seri => case seri of (((ml_compiler, _), filename), _) =>
(fn i => fn seri => case seri of (((ml_compiler, _), (filename, _)), _) =>
let val (l, (out, err)) =
Deep0.Find.build
ml_compiler
@ -1264,12 +1264,13 @@ fun f_command l_mode =
fun mk_fic s = Path.append tmp_export_code (Path.make [s])
val () = Deep0.Find.check_compil ml_compiler ()
val () = Isabelle_System.mkdirs tmp_export_code in
((( (ml_compiler, ml_module)
, Path.implode (if Deep0.Find.export_mode ml_compiler = Deep0.Export_code_env.Directory then
tmp_export_code
else
mk_fic (Deep0.Find.function ml_compiler (Deep0.Find.ext ml_compiler))))
, export_arg), mk_fic)
(( ( (ml_compiler, ml_module)
, ( Path.implode (if Deep0.Find.export_mode ml_compiler = Deep0.Export_code_env.Directory then
tmp_export_code
else
mk_fic (Deep0.Find.function ml_compiler (Deep0.Find.ext ml_compiler)))
, Position.none))
, export_arg), mk_fic)
end)
(List.filter (fn (("self", _), _) => false | _ => true) (#seri_args i_deep))
val _ =
@ -1342,7 +1343,7 @@ let open Generation_mode
:: []))
|> Deep.export_code_cmd' seri_args
tmp_export_code
(fn (((_, _), msg), _) => fn err => if err <> 0 then error msg else ())
(fn (((_, _), (msg, _)), _) => fn err => if err <> 0 then error msg else ())
[name_main]
end
in
@ -1387,7 +1388,7 @@ local
fun fold_thy_shallow f =
META.fold_thy_shallow
(fn f => f () handle ERROR e =>
( warning "Shallow Backtracking: (true) Isabelle declarations occuring among the META-simulated ones are ignored (if any)"
( warning "Shallow Backtracking: (true) Isabelle declarations occurring among the META-simulated ones are ignored (if any)"
(* TODO automatically determine if there is such Isabelle declarations,
for raising earlier a specific error message *)
; error e))

View File

@ -48,10 +48,10 @@ section\<open>Basic Extension of the Standard Library (Depending on RBT)\<close>
theory Init_rbt
imports "../compiler_generic/Init"
"~~/src/HOL/Library/RBT"
"~~/src/HOL/Library/Char_ord"
"~~/src/HOL/Library/List_lexord"
"~~/src/HOL/Library/Product_Lexorder"
"HOL-Library.RBT"
"HOL-Library.Char_ord"
"HOL-Library.List_lexord"
"HOL-Library.Product_Lexorder"
begin
locale RBT

View File

@ -40,7 +40,7 @@
******************************************************************************)
theory Init
imports "~~/src/HOL/Library/Code_Char"
imports "HOL-Library.Code_Char"
"isabelle_home/src/HOL/Isabelle_Main0"
begin

View File

@ -74,8 +74,13 @@ session Isabelle_Meta_Model (AFP) = "HOL-Library" +
"toy_example/embedding/Generator_dynamic_sequential"
"toy_example/generator/Design_deep"
"toy_example/generator/Design_shallow"
"toy_example/document_generated/Design_generated_generated"
"document/Rail"
theories
(* This part ensures that generated theories are accepted:
in general, if X..._generated_generated.thy is wellformed
then we also have X..._generated.thy wellformed *)
"toy_example/document_generated/Design_generated"
"toy_example/document_generated/Design_generated_generated"
document_files
"root.bib"
"root.tex"

View File

@ -56,7 +56,7 @@ Runtime services building on code generation into implementation language SML.
open Basic_Code_Symbol;
(** computation **)
(** ML compiler as evaluation environment **)
(* technical prerequisites *)
@ -72,7 +72,7 @@ fun compile_ML verbose code context =
(* computation as evaluation into ML language values *)
(* evaluation into ML language values *)
@ -82,7 +82,43 @@ fun compile_ML verbose code context =
(** full static evaluation -- still with limited coverage! **)
fun evaluation_code ctxt module_name program tycos consts all_public =
(** generator for computations -- partial implementations of the universal morphism from Isabelle to ML terms **)
(* auxiliary *)
(* possible type signatures of constants *)
(* name mangling *)
(* checks for input terms *)
(* code generation for of the universal morphism *)
(* dedicated preprocessor for computations *)
(* mounting computations *)
(** variants of universal runtime code generation **)
(*FIXME consolidate variants*)
fun runtime_code'' ctxt module_name program tycos consts all_public =
let
val thy = Proof_Context.theory_of ctxt;
val (ml_modules, target_names) =
@ -104,7 +140,7 @@ fun evaluation_code ctxt module_name program tycos consts all_public =
(** code antiquotation **)
(** code and computation antiquotations **)
@ -184,7 +220,7 @@ fun gen_code_reflect prep_type prep_const all_public raw_datatypes raw_functions
val functions = map (prep_const thy) raw_functions;
val consts = constrs @ functions;
val program = Code_Thingol.consts_program ctxt consts;
val result = evaluation_code ctxt module_name program tycos consts all_public
val result = runtime_code'' ctxt module_name program tycos consts all_public
|> (apsnd o apsnd) (chop (length constrs));
in
thy
@ -209,12 +245,12 @@ val _ =
Outer_Syntax.command @{command_keyword code_reflect'}
"enrich runtime environment with generated code"
(Scan.optional (@{keyword "open"} |-- Scan.succeed true) false --
Parse.name -- Scan.optional (@{keyword "datatypes"} |-- Parse.!!! (parse_datatype
Parse.name -- Scan.optional (@{keyword "datatypes"} |-- Parse.!!! (parse_datatype
::: Scan.repeat (@{keyword "and"} |-- parse_datatype))) []
-- Scan.optional (@{keyword "functions"} |-- Parse.!!! (Scan.repeat1 Parse.name)) []
-- Scan.option (@{keyword "file"} |-- Parse.!!! Parse.name)
>> (fn ((((all_public, module_name), raw_datatypes), raw_functions), some_file) => Toplevel.theory
(code_reflect_cmd all_public raw_datatypes raw_functions module_name some_file)))
(code_reflect_cmd all_public raw_datatypes raw_functions module_name some_file)));
end; (*local*)

View File

@ -98,8 +98,14 @@ val parse_inst_ident = Parse.name --| @{keyword "::"} -- Parse.class;
(* code generation *)
fun prep_destination "" = NONE
| prep_destination s = SOME (Path.explode s);
fun prep_destination (s, pos) =
if s = "" then NONE
else
let
val _ = Position.report pos Markup.language_path;
val path = Path.explode s;
val _ = Position.report pos (Markup.path (Path.smart_implode path));
in SOME path end;
fun export_code_cmd all_public raw_cs seris ctxt =

View File

@ -193,7 +193,7 @@ val local_theory_to_proof = local_theory_command Toplevel.local_theory_to_proof;
(* parse commands *)
val bootstrap =
Config.bool (Config.declare ("Outer_Syntax.bootstrap", @{here}) (K (Config.Bool true)));
Config.bool (Config.declare ("Outer_Syntax.bootstrap", \<^here>) (K (Config.Bool true)));
local
@ -340,7 +340,7 @@ fun check_command ctxt (name, pos) =
(* 'ML' command -- required for bootstrapping Isar *)
val _ =
command ("ML", @{here}) "ML text within theory or local theory"
command ("ML", \<^here>) "ML text within theory or local theory"
(Parse.ML_source >> (fn source =>
Toplevel.generic_theory
(ML_Context.exec (fn () =>

View File

@ -79,8 +79,8 @@ sig
val exec_id: Document_ID.exec -> transitions -> transitions
val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
val add_hook: (transition -> state -> state -> unit) -> unit
val get_timing: transition -> Time.time option
val put_timing: Time.time option -> transition -> transition
val get_timing: transition -> Time.time
val put_timing: Time.time -> transition -> transition
val transition: bool -> transition -> state -> state * (exn * string) option
val command_errors: bool -> transition -> state -> Runtime.error list * state option
val command_exception: bool -> transition -> state -> state
@ -296,7 +296,7 @@ datatype state_write = Store_backup | Store_default
datatype transition = Transition of
{name: string, (*command name*)
pos: Position.T, (*source position*)
timing: Time.time option, (*prescient timing information*)
timing: Time.time, (*prescient timing information*)
trans: trans list, (*primitive transitions (union)*)
read_write: state_read * state_write}; (*state update status*)
@ -308,7 +308,7 @@ fun make_transition (name, pos, timing, trans, read_write) =
fun map_transition f (Transition {name, pos, timing, trans, read_write}) =
make_transition (f (name, pos, timing, trans, read_write));
val empty = make_transition ("", Position.none, NONE, [], (Load_previous, Store_default));
val empty = make_transition ("", Position.none, Time.zeroTime, [], (Load_previous, Store_default));
(* diagnostics *)
@ -446,7 +446,7 @@ fun local_theory' restricted target f = present_transaction (fn int =>
fun local_theory restricted target f = local_theory' restricted target (K f);
fun present_local_theory target = present_transaction (fn int =>
fun present_local_theory target = present_transaction (fn _ =>
(fn Theory (gthy, _) =>
let val (finish, lthy) = Named_Target.switch target gthy;
in Theory (finish lthy, SOME lthy) end
@ -581,7 +581,7 @@ fun put_timing timing = map_transition (fn (name, pos, _, trans, read_write) =>
local
fun app int (tr as Transition {name, trans, ...}) =
fun app int (tr as Transition {trans, ...}) =
setmp_thread_position tr (fn state =>
let
val timing_start = Timing.start ();
@ -622,7 +622,7 @@ end;
fun command_errors int tr st =
(case transition int tr st of
(st', NONE) => ([], SOME st')
| (_, SOME (exn, _)) => (Runtime.exn_messages_ids exn, NONE));
| (_, SOME (exn, _)) => (Runtime.exn_messages exn, NONE));
fun command_exception int tr st =
(case transition int tr st of
@ -677,19 +677,9 @@ structure Result = Proof_Data
val get_result = Result.get o Proof.context_of;
val put_result = Proof.map_context o Result.put;
fun timing_estimate include_head elem =
let
val trs = Thy_Syntax.flat_element elem |> not include_head ? tl;
val timings = map get_timing trs;
in
if forall is_some timings then
SOME (fold (curry (op +) o the) timings Time.zeroTime)
else NONE
end;
fun priority NONE = ~1
| priority (SOME estimate) =
Int.min (Real.floor (Real.max (Math.log10 (Time.toReal estimate), ~3.0)) - 3, ~1);
fun timing_estimate elem =
let val trs = tl (Thy_Syntax.flat_element elem)
in fold (fn tr => fn t => get_timing tr + t) trs Time.zeroTime end;
fun proof_future_enabled estimate st =
(case try proof_of st of
@ -698,18 +688,14 @@ fun proof_future_enabled estimate st =
not (Proof.is_relevant state) andalso
(if can (Proof.assert_bottom true) state
then Goal.future_enabled 1
else
(case estimate of
NONE => Goal.future_enabled 2
| SOME t => Goal.future_enabled_timing t)));
else Goal.future_enabled 2 orelse Goal.future_enabled_timing estimate));
fun atom_result keywords tr st =
let
val st' =
if Goal.future_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then
(Execution.fork
{name = "Toplevel.diag", pos = pos_of tr,
pri = priority (timing_estimate true (Thy_Syntax.atom tr))}
{name = "Toplevel.diag", pos = pos_of tr, pri = ~1}
(fn () => command tr st); st)
else command tr st;
in (Result (tr, st'), st') end;
@ -721,7 +707,7 @@ fun element_result keywords (Thy_Syntax.Element (tr, NONE)) st = atom_result key
let
val (head_result, st') = atom_result keywords head_tr st;
val (body_elems, end_tr) = element_rest;
val estimate = timing_estimate false elem;
val estimate = timing_estimate elem;
in
if not (proof_future_enabled estimate st')
then
@ -736,7 +722,7 @@ fun element_result keywords (Thy_Syntax.Element (tr, NONE)) st = atom_result key
val future_proof =
Proof.future_proof (fn state =>
Execution.fork
{name = "Toplevel.future_proof", pos = pos_of head_tr, pri = priority estimate}
{name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
(fn () =>
let
val State (SOME (Proof (prf, (_, orig_gthy))), prev, backup) = st';

View File

@ -12,6 +12,7 @@ sig
val read: Keyword.keywords -> theory -> Path.T-> (unit -> theory) ->
blob list * int -> Token.T list -> Toplevel.transitions
type eval
val eval_exec_id: eval -> Document_ID.exec
val eval_eq: eval * eval -> bool
val eval_running: eval -> bool
val eval_finished: eval -> bool
@ -116,7 +117,7 @@ fun reports_of_token keywords tok =
Input.source_explode (Token.input_of tok)
|> map_filter (fn (sym, pos) =>
if Symbol.is_malformed sym
then SOME ((pos, Markup.bad), "Malformed symbolic character") else NONE);
then SOME ((pos, Markup.bad ()), "Malformed symbolic character") else NONE);
val is_malformed = Token.is_error tok orelse not (null malformed_symbols);
val reports = Token.reports keywords tok @ Token.completion_report tok @ malformed_symbols;
in (is_malformed, reports) end;
@ -201,7 +202,7 @@ fun check_cmts span tr st' =
(Thy_Output.output_text st' {markdown = false} (Token.input_of cmt); [])
handle exn =>
if Exn.is_interrupt exn then Exn.reraise exn
else Runtime.exn_messages_ids exn)) ();
else Runtime.exn_messages exn)) ();
fun report tr m =
Toplevel.setmp_thread_position tr (fn () => Output.report [Markup.markup_only m]) ();
@ -242,7 +243,7 @@ fun eval_state' keywords span tr state =
let
val _ = status tr Markup.failed;
val _ = status tr Markup.finished;
val _ = if null errs then (report tr Markup.bad; Exn.interrupt ()) else ();
val _ = if null errs then (report tr (Markup.bad ()); Exn.interrupt ()) else ();
in {failed = true, command = tr, state = st} end
| SOME st' =>
let
@ -274,7 +275,7 @@ fun eval keywords master_dir init blobs_info span eval0 =
(fn () =>
read keywords thy master_dir init blobs_info span |> Toplevel.exec_id exec_id) ();
in eval_state keywords span tr eval_state0 end;
in Eval {exec_id = exec_id, eval_process = Lazy.lazy process} end;
in Eval {exec_id = exec_id, eval_process = Lazy.lazy_name "Command.eval" process} end;
end;
@ -303,7 +304,7 @@ fun print_error tr opt_context e =
(Toplevel.setmp_thread_position tr o Runtime.controlled_execution opt_context) e ()
handle exn =>
if Exn.is_interrupt exn then Exn.reraise exn
else List.app (Future.error_message (Toplevel.pos_of tr)) (Runtime.exn_messages_ids exn);
else List.app (Future.error_message (Toplevel.pos_of tr)) (Runtime.exn_messages exn);
fun print_finished (Print {print_process, ...}) = Lazy.is_finished print_process;
@ -332,7 +333,7 @@ fun print command_visible command_overlays keywords command_name eval old_prints
in
Print {
name = name, args = args, delay = delay, pri = pri, persistent = persistent,
exec_id = exec_id, print_process = Lazy.lazy process}
exec_id = exec_id, print_process = Lazy.lazy_name "Command.print" process}
end;
fun bad_print name args exn =

View File

@ -1,283 +0,0 @@
(* Title: Pure/PIDE/resources.ML
Author: Makarius
Resources for theories and auxiliary files.
*)
signature RESOURCES =
sig
val master_directory: theory -> Path.T
val imports_of: theory -> (string * Position.T) list
val thy_path: Path.T -> Path.T
val check_thy: Path.T -> string ->
{master: Path.T * SHA1.digest, text: string, theory_pos: Position.T,
imports: (string * Position.T) list, keywords: Thy_Header.keywords}
val parse_files: string -> (theory -> Token.file list) parser
val provide: Path.T * SHA1.digest -> theory -> theory
val provide_parse_files: string -> (theory -> Token.file list * theory) parser
val loaded_files_current: theory -> bool
val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory
val load_thy: bool -> HTML.symbols -> (Toplevel.transition -> Time.time option) -> int ->
Path.T -> Thy_Header.header -> Position.T -> string -> theory list ->
theory * (unit -> unit) * int
end;
structure Resources: RESOURCES =
struct
(* manage source files *)
type files =
{master_dir: Path.T, (*master directory of theory source*)
imports: (string * Position.T) list, (*source specification of imports*)
provided: (Path.T * SHA1.digest) list}; (*source path, digest*)
fun make_files (master_dir, imports, provided): files =
{master_dir = master_dir, imports = imports, provided = provided};
structure Files = Theory_Data
(
type T = files;
val empty = make_files (Path.current, [], []);
fun extend _ = empty;
fun merge _ = empty;
);
fun map_files f =
Files.map (fn {master_dir, imports, provided} =>
make_files (f (master_dir, imports, provided)));
val master_directory = #master_dir o Files.get;
val imports_of = #imports o Files.get;
fun put_deps master_dir imports = map_files (fn _ => (master_dir, imports, []));
(* theory files *)
val thy_path = Path.ext "thy";
fun check_file dir file = File.check_file (File.full_path dir file);
fun check_thy dir thy_name =
let
val path = thy_path (Path.basic thy_name);
val master_file = check_file dir path;
val text = File.read master_file;
val {name = (name, pos), imports, keywords} =
Thy_Header.read (Path.position master_file) text;
val _ = thy_name <> name andalso
error ("Bad theory name " ^ quote name ^
" for file " ^ Path.print path ^ Position.here pos);
in
{master = (master_file, SHA1.digest text), text = text, theory_pos = pos,
imports = imports, keywords = keywords}
end;
(* load files *)
fun parse_files cmd =
Scan.ahead Parse.not_eof -- Parse.path >> (fn (tok, name) => fn thy =>
(case Token.get_files tok of
[] =>
let
val keywords = Thy_Header.get_keywords thy;
val master_dir = master_directory thy;
val pos = Token.pos_of tok;
val src_paths = Keyword.command_files keywords cmd (Path.explode name);
in map (Command.read_file master_dir pos) src_paths end
| files => map Exn.release files));
fun provide (src_path, id) =
map_files (fn (master_dir, imports, provided) =>
if AList.defined (op =) provided src_path then
error ("Duplicate use of source file: " ^ Path.print src_path)
else (master_dir, imports, (src_path, id) :: provided));
fun provide_parse_files cmd =
parse_files cmd >> (fn files => fn thy =>
let
val fs = files thy;
val thy' = fold (fn {src_path, digest, ...} => provide (src_path, digest)) fs thy;
in (fs, thy') end);
fun load_file thy src_path =
let
val full_path = check_file (master_directory thy) src_path;
val text = File.read full_path;
val id = SHA1.digest text;
in ((full_path, id), text) end;
fun loaded_files_current thy =
#provided (Files.get thy) |>
forall (fn (src_path, id) =>
(case try (load_file thy) src_path of
NONE => false
| SOME ((_, id'), _) => id = id'));
(* load theory *)
fun begin_theory master_dir {name, imports, keywords} parents =
Theory.begin_theory name parents
|> put_deps master_dir imports
|> Thy_Header.add_keywords keywords;
datatype span_raw =
Span_cmd of Command_Span.span
| Span_tr of Toplevel.transition
fun excursion keywords master_dir last_timing init elements =
let
fun prepare_span st = fn
Span_cmd span =>
Command_Span.content span
|> Command.read keywords (Command.read_thy st) master_dir init ([], ~1)
|> map (fn tr => Toplevel.put_timing (last_timing tr) tr)
| Span_tr tr => [tr];
fun element_result elem (st, _) =
let
val (results, st') = Toplevel.element_result keywords elem st;
val pos' = Toplevel.pos_of (Thy_Syntax.last_element elem);
in (results, (st', pos')) end;
val meta_cmd = fn [_] => false | _ => true
val (results, (end_state, end_pos)) =
let fun aux _ ([], acc) = ([], acc)
| aux prev_xs ((x :: xs), acc) =
let
val x0 = Thy_Syntax.map_element (prepare_span (#1 acc)) x
in
if Thy_Syntax.exists_element meta_cmd x0 then
let val (l_reparse, prev_xs) =
if case x0 of Thy_Syntax.Element (a0, _) => meta_cmd a0 then
prev_xs |>
(Scan.permissive (Scan.one (fn (Thy_Syntax.Element (_, NONE), _) => true | _ => false) >> (fn l => [l]))
|| Scan.succeed [])
else ([], prev_xs)
in
aux
prev_xs
(apfst
(fn l =>
Thy_Syntax.parse_elements keywords
(fn x => Span_cmd (Command_Span.Span (x, [])))
(fn Span_cmd (Command_Span.Span (x, _)) => x
| Span_tr tr => Command_Span.Command_Span (Toplevel.name_of tr, Toplevel.pos_of tr))
(List.concat (List.concat [ l,
map (map Span_tr) (Thy_Syntax.flat_element x0),
map Thy_Syntax.flat_element xs])))
(case map (apfst Thy_Syntax.flat_element) (rev l_reparse) of
[] => ([], acc)
| (x, acc) :: xs => (x :: map #1 xs, acc)))
end
else
let
val x0 = Thy_Syntax.map_element hd x0
val (x', acc') = element_result x0 acc;
val (xs', acc'') = aux ((Thy_Syntax.map_element Span_tr x0, acc) :: prev_xs) (xs, acc');
in (x' :: xs', acc'') end
end
in
aux [] (elements, (Toplevel.toplevel, Position.none))
end;
val thy = Toplevel.end_theory end_pos end_state;
in (results, thy) end;
fun load_thy document symbols last_timing update_time master_dir header text_pos text parents =
let
val (name, _) = #name header;
val keywords =
fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents
(Keyword.add_keywords (#keywords header) Keyword.empty_keywords);
val toks = Token.explode keywords text_pos text;
val spans = Outer_Syntax.parse_spans toks;
val elements = map (Thy_Syntax.map_element Span_cmd)
(Thy_Syntax.parse_elements keywords
(fn x => Command_Span.Span (x, []))
(fn Command_Span.Span (x, _) => x)
spans)
fun init () =
begin_theory master_dir header parents
|> Present.begin_theory update_time
(fn () => implode (map (HTML.present_span symbols keywords) spans));
val (results, thy) =
cond_timeit true ("theory " ^ quote name)
(fn () => excursion keywords master_dir last_timing init elements);
fun present () =
let
val res = filter_out (Toplevel.is_ignored o #1) (maps Toplevel.join_results results);
in
if exists (Toplevel.is_skipped_proof o #2) res then
warning ("Cannot present theory with skipped proofs: " ^ quote name)
else
if document then let val tex_source = Thy_Output.present_thy thy res toks |> Buffer.content in Present.theory_output name tex_source end else ()
end;
in (thy, present, size text) end;
(* antiquotations *)
local
fun err msg pos = error (msg ^ Position.here pos);
fun check_path check_file ctxt dir (name, pos) =
let
val _ = Context_Position.report ctxt pos Markup.language_path;
val path = Path.append dir (Path.explode name) handle ERROR msg => err msg pos;
val _ = Path.expand path handle ERROR msg => err msg pos;
val _ = Context_Position.report ctxt pos (Markup.path (Path.smart_implode path));
val _ =
(case check_file of
NONE => path
| SOME check => (check path handle ERROR msg => err msg pos));
in path end;
fun document_antiq check_file ctxt (name, pos) =
let
val dir = master_directory (Proof_Context.theory_of ctxt);
val _ = check_path check_file ctxt dir (name, pos);
in
space_explode "/" name
|> map Latex.output_ascii
|> space_implode (Latex.output_ascii "/" ^ "\\discretionary{}{}{}")
|> enclose "\\isatt{" "}"
end;
fun ML_antiq check_file ctxt (name, pos) =
let val path = check_path check_file ctxt Path.current (name, pos);
in "Path.explode " ^ ML_Syntax.print_string (Path.implode path) end;
in
val _ = Theory.setup
(Thy_Output.antiquotation @{binding path} (Scan.lift (Parse.position Parse.path))
(document_antiq NONE o #context) #>
Thy_Output.antiquotation @{binding file} (Scan.lift (Parse.position Parse.path))
(document_antiq (SOME File.check_file) o #context) #>
Thy_Output.antiquotation @{binding dir} (Scan.lift (Parse.position Parse.path))
(document_antiq (SOME File.check_dir) o #context) #>
ML_Antiquotation.value @{binding path}
(Args.context -- Scan.lift (Parse.position Parse.path)
>> uncurry (ML_antiq NONE)) #>
ML_Antiquotation.value @{binding file}
(Args.context -- Scan.lift (Parse.position Parse.path)
>> uncurry (ML_antiq (SOME File.check_file))) #>
ML_Antiquotation.value @{binding dir}
(Args.context -- Scan.lift (Parse.position Parse.path)
>> uncurry (ML_antiq (SOME File.check_dir))));
end;
end;

View File

@ -16,10 +16,10 @@ sig
val use_theories:
{document: bool,
symbols: HTML.symbols,
last_timing: Toplevel.transition -> Time.time option,
last_timing: Toplevel.transition -> Time.time,
qualifier: string,
master_dir: Path.T} -> (string * Position.T) list -> unit
val use_thys: (string * Position.T) list -> unit
val use_thy: string * Position.T -> unit
val use_thy: string -> unit
val script_thy: Position.T -> string -> theory -> theory
val register_thy: theory -> unit
val finish: unit -> unit
@ -54,11 +54,9 @@ type deps =
fun make_deps master imports : deps = {master = master, imports = imports};
fun master_dir (d: deps option) =
fun master_dir_deps (d: deps option) =
the_default Path.current (Option.map (Path.dir o #1 o #master) d);
fun base_name s = Path.implode (Path.base (Path.explode s));
local
val global_thys =
Synchronized.var "Thy_Info.thys"
@ -88,7 +86,7 @@ fun get_thy name = get (get_thys ()) name;
val lookup_deps = Option.map #1 o lookup_thy;
val master_directory = master_dir o #1 o get_thy;
val master_directory = master_dir_deps o #1 o get_thy;
(* access theory *)
@ -130,8 +128,8 @@ val remove_thy = change_thys o remove;
fun update deps theory thys =
let
val name = Context.theory_name theory;
val parents = map Context.theory_name (Theory.parents_of theory);
val name = Context.theory_long_name theory;
val parents = map Context.theory_long_name (Theory.parents_of theory);
val thys' = remove name thys;
val _ = map (get thys') parents;
@ -156,11 +154,10 @@ fun result_ord (Result {weight = i, ...}, Result {weight = j, ...}) = int_ord (j
fun join_theory (Result {theory, exec_id, ...}) =
let
(*toplevel proofs and diags*)
val _ = Future.join_tasks (maps Future.group_snapshot (Execution.peek exec_id));
(*fully nested proofs*)
val res = Exn.capture Thm.join_theory_proofs theory;
in res :: map Exn.Exn (maps Task_Queue.group_status (Execution.peek exec_id)) end;
val _ = Execution.join [exec_id];
val res = Exn.capture Thm.consolidate_theory theory;
val exns = maps Task_Queue.group_status (Execution.peek exec_id);
in res :: map Exn.Exn exns end;
datatype task =
Task of Path.T * string list * (theory list -> result) |
@ -171,8 +168,6 @@ fun task_finished (Task _) = false
fun task_parents deps (parents: string list) = map (the o AList.lookup (op =) deps) parents;
local
val schedule_seq =
String_Graph.schedule (fn deps => fn (_, task) =>
(case task of
@ -224,12 +219,106 @@ val schedule_futures = Thread_Attributes.uninterruptible (fn _ => fn tasks =>
val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4);
in () end);
in
fun schedule_tasks tasks =
if Multithreading.enabled () then schedule_futures tasks else schedule_seq tasks;
(* eval theory *)
end;
datatype span_raw =
Span_cmd of Command_Span.span
| Span_tr of Toplevel.transition
fun excursion keywords master_dir last_timing init elements =
let
fun prepare_span st = fn
Span_cmd span =>
Command_Span.content span
|> Command.read keywords (Command.read_thy st) master_dir init ([], ~1)
|> map (fn tr => Toplevel.put_timing (last_timing tr) tr)
| Span_tr tr => [tr];
fun element_result elem (st, _) =
let
val (results, st') = Toplevel.element_result keywords elem st;
val pos' = Toplevel.pos_of (Thy_Syntax.last_element elem);
in (results, (st', pos')) end;
val meta_cmd = fn [_] => false | _ => true
val (results, (end_state, end_pos)) =
let fun aux _ ([], acc) = ([], acc)
| aux prev_xs ((x :: xs), acc) =
let
val x0 = Thy_Syntax.map_element (prepare_span (#1 acc)) x
in
if Thy_Syntax.exists_element meta_cmd x0 then
let val (l_reparse, prev_xs) =
if case x0 of Thy_Syntax.Element (a0, _) => meta_cmd a0 then
prev_xs |>
(Scan.permissive (Scan.one (fn (Thy_Syntax.Element (_, NONE), _) => true | _ => false) >> (fn l => [l]))
|| Scan.succeed [])
else ([], prev_xs)
in
aux
prev_xs
(apfst
(fn l =>
Thy_Syntax.parse_elements keywords
(fn x => Span_cmd (Command_Span.Span (x, [])))
(fn Span_cmd (Command_Span.Span (x, _)) => x
| Span_tr tr => Command_Span.Command_Span (Toplevel.name_of tr, Toplevel.pos_of tr))
(List.concat (List.concat [ l,
map (map Span_tr) (Thy_Syntax.flat_element x0),
map Thy_Syntax.flat_element xs])))
(case map (apfst Thy_Syntax.flat_element) (rev l_reparse) of
[] => ([], acc)
| (x, acc) :: xs => (x :: map #1 xs, acc)))
end
else
let
val x0 = Thy_Syntax.map_element hd x0
val (x', acc') = element_result x0 acc;
val (xs', acc'') = aux ((Thy_Syntax.map_element Span_tr x0, acc) :: prev_xs) (xs, acc');
in (x' :: xs', acc'') end
end
in
aux [] (elements, (Toplevel.toplevel, Position.none))
end;
val thy = Toplevel.end_theory end_pos end_state;
in (results, thy) end;
fun eval_thy document symbols last_timing update_time master_dir header text_pos text parents =
let
val (name, _) = #name header;
val keywords =
fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents
(Keyword.add_keywords (#keywords header) Keyword.empty_keywords);
val toks = Token.explode keywords text_pos text;
val spans = Outer_Syntax.parse_spans toks;
val elements = map (Thy_Syntax.map_element Span_cmd)
(Thy_Syntax.parse_elements keywords
(fn x => Command_Span.Span (x, []))
(fn Command_Span.Span (x, _) => x)
spans)
fun init () =
Resources.begin_theory master_dir header parents
|> Present.begin_theory update_time
(fn () => implode (map (HTML.present_span symbols keywords) spans));
val (results, thy) =
cond_timeit true ("theory " ^ quote name)
(fn () => excursion keywords master_dir last_timing init elements);
fun present () =
let
val res = filter_out (Toplevel.is_ignored o #1) (maps Toplevel.join_results results);
in
if exists (Toplevel.is_skipped_proof o #2) res then
warning ("Cannot present theory with skipped proofs: " ^ quote name)
else
if document then let val tex_source = Thy_Output.present_thy thy res toks |> Buffer.content in Present.theory_output thy tex_source end else ()
end;
in (thy, present, size text) end;
(* require_thy -- checking database entries wrt. the file-system *)
@ -259,7 +348,7 @@ fun load_thy document symbols last_timing
val text_pos = Position.put_id (Document_ID.print exec_id) (Path.position thy_path);
val (theory, present, weight) =
Resources.load_thy document symbols last_timing update_time dir header text_pos text
eval_thy document symbols last_timing update_time dir header text_pos text
(if name = Context.PureN then [Context.the_global_context ()] else parents);
fun commit () = update_thy deps theory;
in
@ -286,44 +375,45 @@ fun check_deps dir name =
in
fun require_thys document symbols last_timing initiators dir strs tasks =
fold_map (require_thy document symbols last_timing initiators dir) strs tasks |>> forall I
and require_thy document symbols last_timing initiators dir (str, require_pos) tasks =
fun require_thys document symbols last_timing initiators qualifier dir strs tasks =
fold_map (require_thy document symbols last_timing initiators qualifier dir) strs tasks
|>> forall I
and require_thy document symbols last_timing initiators qualifier dir (s, require_pos) tasks =
let
val path = Path.expand (Path.explode str);
val name = Path.implode (Path.base path);
val node_name = File.full_path dir (Resources.thy_path path);
val {node_name, master_dir, theory_name} = Resources.import_name qualifier dir s;
fun check_entry (Task (node_name', _, _)) =
if op = (apply2 File.platform_path (node_name, node_name'))
then ()
else
error ("Incoherent imports for theory " ^ quote name ^
error ("Incoherent imports for theory " ^ quote theory_name ^
Position.here require_pos ^ ":\n" ^
" " ^ Path.print node_name ^ "\n" ^
" " ^ Path.print node_name')
| check_entry _ = ();
in
(case try (String_Graph.get_node tasks) name of
(case try (String_Graph.get_node tasks) theory_name of
SOME task => (check_entry task; (task_finished task, tasks))
| NONE =>
let
val dir' = Path.append dir (Path.dir path);
val _ = member (op =) initiators name andalso error (cycle_msg initiators);
val _ = member (op =) initiators theory_name andalso error (cycle_msg initiators);
val (current, deps, theory_pos, imports, keywords) = check_deps dir' name
val (current, deps, theory_pos, imports, keywords) = check_deps master_dir theory_name
handle ERROR msg =>
cat_error msg
("The error(s) above occurred for theory " ^ quote name ^
("The error(s) above occurred for theory " ^ quote theory_name ^
Position.here require_pos ^ required_by "\n" initiators);
val parents = map (base_name o #1) imports;
val qualifier' = Resources.theory_qualifier theory_name;
val dir' = Path.append dir (master_dir_deps (Option.map #1 deps));
val parents = map (#theory_name o Resources.import_name qualifier' dir' o #1) imports;
val (parents_current, tasks') =
require_thys document symbols last_timing (name :: initiators)
(Path.append dir (master_dir (Option.map #1 deps))) imports tasks;
require_thys document symbols last_timing (theory_name :: initiators)
qualifier' dir' imports tasks;
val all_current = current andalso parents_current;
val task =
if all_current then Finished (get_theory name)
if all_current then Finished (get_theory theory_name)
else
(case deps of
NONE => raise Fail "Malformed deps"
@ -332,27 +422,29 @@ and require_thy document symbols last_timing initiators dir (str, require_pos) t
val update_time = serial ();
val load =
load_thy document symbols last_timing initiators update_time dep
text (name, theory_pos) keywords;
text (theory_name, theory_pos) keywords;
in Task (node_name, parents, load) end);
val tasks'' = new_entry name parents task tasks';
val tasks'' = new_entry theory_name parents task tasks';
in (all_current, tasks'') end)
end;
end;
(* use_thy *)
(* use theories *)
fun use_theories {document, symbols, last_timing, master_dir} imports =
schedule_tasks
(snd (require_thys document symbols last_timing [] master_dir imports String_Graph.empty));
fun use_theories {document, symbols, last_timing, qualifier, master_dir} imports =
let
val (_, tasks) =
require_thys document symbols last_timing [] qualifier master_dir imports String_Graph.empty;
in if Multithreading.enabled () then schedule_futures tasks else schedule_seq tasks end;
val use_thys =
fun use_thy name =
use_theories
{document = false, symbols = HTML.no_symbols, last_timing = K NONE, master_dir = Path.current};
val use_thy = use_thys o single;
{document = false, symbols = HTML.no_symbols, last_timing = K Time.zeroTime,
qualifier = Resources.default_qualifier, master_dir = Path.current}
[(name, Position.none)];
(* toplevel scripting -- without maintaining database *)
@ -371,7 +463,7 @@ fun script_thy pos txt thy =
fun register_thy theory =
let
val name = Context.theory_name theory;
val name = Context.theory_long_name theory;
val {master, ...} = Resources.check_thy (Resources.master_directory theory) name;
val imports = Resources.imports_of theory;
in
@ -389,4 +481,4 @@ fun finish () = change_thys (String_Graph.map (fn _ => fn (_, entry) => (NONE, e
end;
fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy (name, Position.none));
fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy name);

View File

@ -44,13 +44,13 @@ struct
(** options **)
val display = Attrib.setup_option_bool ("thy_output_display", @{here});
val break = Attrib.setup_option_bool ("thy_output_break", @{here});
val quotes = Attrib.setup_option_bool ("thy_output_quotes", @{here});
val margin = Attrib.setup_option_int ("thy_output_margin", @{here});
val indent = Attrib.setup_option_int ("thy_output_indent", @{here});
val source = Attrib.setup_option_bool ("thy_output_source", @{here});
val modes = Attrib.setup_option_string ("thy_output_modes", @{here});
val display = Attrib.setup_option_bool ("thy_output_display", \<^here>);
val break = Attrib.setup_option_bool ("thy_output_break", \<^here>);
val quotes = Attrib.setup_option_bool ("thy_output_quotes", \<^here>);
val margin = Attrib.setup_option_int ("thy_output_margin", \<^here>);
val indent = Attrib.setup_option_int ("thy_output_indent", \<^here>);
val source = Attrib.setup_option_bool ("thy_output_source", \<^here>);
val modes = Attrib.setup_option_string ("thy_output_modes", \<^here>);
structure Wrappers = Proof_Data
@ -169,7 +169,7 @@ fun eval_antiq state (opts, src) =
fun cmd ctxt = wrap ctxt (fn () => command src state ctxt) ();
val _ = cmd preview_ctxt;
val print_modes = space_explode "," (Config.get print_ctxt modes) @ Latex.modes;
val print_modes = space_explode "," (Config.get print_ctxt modes) @ [Latex.latexN];
in Print_Mode.with_modes print_modes (fn () => cmd print_ctxt) () end;
in
@ -591,7 +591,7 @@ fun pretty_type ctxt s =
let val Type (name, _) = Proof_Context.read_type_name {proper = true, strict = false} ctxt s
in Pretty.str (Proof_Context.extern_type ctxt name) end;
fun pretty_prf full ctxt = Proof_Syntax.pretty_proof_of ctxt full;
fun pretty_prf full ctxt = Proof_Syntax.pretty_clean_proof_of ctxt full;
fun pretty_theory ctxt (name, pos) = (Theory.check ctxt (name, pos); Pretty.str name);

View File

@ -219,7 +219,7 @@ fun check l_oid l =
META.check_export_code
(writeln o Mi)
(warning o Mi)
(writeln o Markup.markup Markup.bad o Mi)
(fn s => writeln (Markup.markup (Markup.bad ()) (Mi s)))
(error o To_string0)
(Ml (Mp I Me) l_oid)
((META.SS_base o META.ST) l)
@ -1061,7 +1061,7 @@ fun export_code_tmp_file seris g =
else
Isabelle_System.with_tmp_file tmp_name (Deep0.Find.ext ml_compiler))
(fn filename =>
g (((((ml_compiler, ml_module), Path.implode filename), export_arg) :: accu)))
g (((((ml_compiler, ml_module), (Path.implode filename, Position.none)), export_arg) :: accu)))
end))
seris
(fn f => f [])
@ -1082,7 +1082,7 @@ fun export_code_cmd' seris tmp_export_code f_err raw_cs thy =
let val v = Deep0.apply_hs_code_identifiers Deep0.Export_code_env.Haskell.argument thy in
if mem_scala then Code_printing.apply_code_printing v else v end) in
List_mapi
(fn i => fn seri => case seri of (((ml_compiler, _), filename), _) =>
(fn i => fn seri => case seri of (((ml_compiler, _), (filename, _)), _) =>
let val (l, (out, err)) =
Deep0.Find.build
ml_compiler
@ -1259,12 +1259,13 @@ fun f_command l_mode =
fun mk_fic s = Path.append tmp_export_code (Path.make [s])
val () = Deep0.Find.check_compil ml_compiler ()
val () = Isabelle_System.mkdirs tmp_export_code in
((( (ml_compiler, ml_module)
, Path.implode (if Deep0.Find.export_mode ml_compiler = Deep0.Export_code_env.Directory then
tmp_export_code
else
mk_fic (Deep0.Find.function ml_compiler (Deep0.Find.ext ml_compiler))))
, export_arg), mk_fic)
(( ( (ml_compiler, ml_module)
, ( Path.implode (if Deep0.Find.export_mode ml_compiler = Deep0.Export_code_env.Directory then
tmp_export_code
else
mk_fic (Deep0.Find.function ml_compiler (Deep0.Find.ext ml_compiler)))
, Position.none))
, export_arg), mk_fic)
end)
(List.filter (fn (("self", _), _) => false | _ => true) (#seri_args i_deep))
val _ =
@ -1337,7 +1338,7 @@ let open Generation_mode
:: []))
|> Deep.export_code_cmd' seri_args
tmp_export_code
(fn (((_, _), msg), _) => fn err => if err <> 0 then error msg else ())
(fn (((_, _), (msg, _)), _) => fn err => if err <> 0 then error msg else ())
[name_main]
end
in
@ -1382,7 +1383,7 @@ local
fun fold_thy_shallow f =
META.fold_thy_shallow
(fn f => f () handle ERROR e =>
( warning "Shallow Backtracking: (true) Isabelle declarations occuring among the META-simulated ones are ignored (if any)"
( warning "Shallow Backtracking: (true) Isabelle declarations occurring among the META-simulated ones are ignored (if any)"
(* TODO automatically determine if there is such Isabelle declarations,
for raising earlier a specific error message *)
; error e))

View File

@ -43,10 +43,10 @@ section\<open>Basic Extension of the Standard Library (Depending on RBT)\<close>
theory Init_rbt
imports "../../Init"
"~~/src/HOL/Library/RBT"
"~~/src/HOL/Library/Char_ord"
"~~/src/HOL/Library/List_lexord"
"~~/src/HOL/Library/Product_Lexorder"
"HOL-Library.RBT"
"HOL-Library.Char_ord"
"HOL-Library.List_lexord"
"HOL-Library.Product_Lexorder"
begin
locale RBT

View File

@ -401,7 +401,7 @@
morekeywords={text,txt,finally,next,also,with,moreover,ultimately,thus,prefer,defer,declare,apply,of,OF,THEN,intros,in,fix,assume,from,this,show,have,and,note,let,hence,where,using},% then, and
classoffset=2,%
keywordstyle=\color{Blue}\textbf,%
morekeywords={axclass,class,instance,recdef,primrec,constdefs,consts_code,types_code,consts,axioms,syntax,typedecl,arities,types,translations,inductive,typedef,datatype,record,instance,defs,specification,proof,test_spec,lemmas,lemma,assumes,shows,definition,fun,function,theorem,case},%
morekeywords={axclass,class,instance,primrec,constdefs,consts_code,types_code,consts,axioms,syntax,typedecl,arities,types,translations,inductive,typedef,datatype,record,instance,defs,specification,proof,test_spec,lemmas,lemma,assumes,shows,definition,fun,function,theorem,case},%
classoffset=3,%
keywordstyle=\color{BrickRed}\textbf,%
morekeywords={oops,sorry},%