Fix thm and file anti-quotations short name bug
ci/woodpecker/push/build Pipeline was successful Details

This commit is contained in:
Nicolas Méric 2023-03-13 10:27:31 +01:00
parent d2a1808fa8
commit 37afd975b3
4 changed files with 22 additions and 36 deletions

View File

@ -55,27 +55,25 @@ print_doc_items
section\<open>Term-Antiquotations Referring to \<^verbatim>\<open>thm\<close>s\<close>
text\<open>Some sample lemma:\<close>
lemma*[l::E] murks : "Example = @{thm ''HOL.refl''}" oops
lemma*[l::E] murks : "Example = @{thm ''refl''}" oops
text-assert-error\<open>... @{E "l"}\<close>\<open>Undefined instance:\<close> \<comment> \<open>oops retracts the ENTIRE system state,
thus also the creation of an instance of E\<close>
lemma*[l::E] local_sample_lemma :
"@{thm \<open>HOL.refl\<close>} = @{thm ''HOL.refl''}" by simp
"@{thm \<open>refl\<close>} = @{thm ''refl''}" by simp
\<comment> \<open>un-evaluated references are similar to
uninterpreted constants. Not much is known
about them, but that doesn't mean that we
can't prove some basics over them...\<close>
(* BUG: Why does this not work ? Shortnames are apparently not recognized !*)
(* SHOULD WORK :
lemma*[l2::E] local_sample_lemma :
lemma*[l2::E] local_sample_lemma2 :
"@{thm ''local_sample_lemma''} = @{thm ''local_sample_lemma''}" by simp
*)
value*\<open>@{thm ''Concept_TermAntiquotations.local_sample_lemma''}\<close>
value-assert-error\<open> @{thm ''Conept_TermAntiquotations.local_sample_lemma''}\<close>\<open>No Theorem\<close>
value*\<open>@{thm ''local_sample_lemma''}\<close>
value-assert-error\<open> @{thm \<open>Conxept_TermAntiquotations.local_sample_lemma\<close>}\<close>\<open>Undefined fact\<close>
section\<open>Testing the Standard ("Built-in") Term-Antiquotations\<close>
@ -104,13 +102,11 @@ section\<open>Other Built-In Term Antiquotations\<close>
text-assert-error[ae::text_element]\<open>@{file "non-existing.thy"}\<close>\<open>No such file: \<close>
text\<open>A text-antiquotation from Main: @{file "TestKit.thy"}\<close>
(* BUG : This shoulöd work rather than throw an exception *)
value-assert-error\<open>@{file \<open>TestKit.thy\<close>}\<close>\<open>No such file: \<close>
text-assert-error\<open>A text-antiquotation from Main: @{file "TestKit.thy"}\<close>
value-assert-error\<open>@{file \<open>non-existing.thy\<close>}\<close>\<open>No such file: \<close>
value*\<open>@{file \<open>TestKit.thy\<close>}\<close>
(* BUG. Again: Long-name Short Name on Files ?
text*[xcv::F, u="@{file ''TestKit.thy''}"]\<open>Lorem ipsum ...\<close>
*)
value*\<open>@{term \<open>aa + bb\<close>}\<close>
value*\<open>@{typ \<open>'a list\<close>}\<close>

View File

@ -5,7 +5,7 @@ text\<open>Term Annotation Antiquotations (TA) can be evaluated with the help of
theory
Evaluation
imports
"Isabelle_DOF-Unit-Tests.TermAntiquotations"
"Isabelle_DOF-Unit-Tests.Concept_TermAntiquotations"
"Isabelle_DOF-Unit-Tests.Concept_High_Level_Invariants"
TestKit
begin

View File

@ -124,10 +124,8 @@ fun error_match2 [_, src] msg = error_match src msg
val _ =
Outer_Syntax.command ("text-assert-error", @{here}) "formal comment macro"
(ODL_Meta_Args_Parser.opt_attributes -- Parse.opt_target -- Scan.repeat1 Parse.document_source
>> (Toplevel.theory o
(fn ((meta_args, xstring_opt), source) =>
(gen_enriched_document_command3 error_match2 "TTT" {body=true}
I I {markdown = true} ((meta_args, xstring_opt), source)))));
>> (Toplevel.theory o (gen_enriched_document_command3 error_match2 "TTT" {body=true}
I I {markdown = true} )));
fun update_instance_command (args,src) thy =
(Monitor_Command_Parser.update_instance_command args thy

View File

@ -487,8 +487,8 @@ struct
datatype monitor_info = Monitor_Info of
{accepted_cids : string list,
rejected_cids : string list,
{accepted_cids : RegExpInterface.env,
rejected_cids : RegExpInterface.env,
automatas : RegExpInterface.automaton list}
@ -753,14 +753,11 @@ fun check_invs get_ml_invs cid_long oid is_monitor thy =
|> map (fn check => check oid is_monitor ctxt)
in (fold_and checks) end)
fun check_ml_invs cid_long oid is_monitor thy =
check_invs get_ml_invariants cid_long oid is_monitor thy
val check_ml_invs = check_invs get_ml_invariants
fun check_opening_ml_invs cid_long oid is_monitor thy =
check_invs get_opening_ml_invariants cid_long oid is_monitor thy
val check_opening_ml_invs = check_invs get_opening_ml_invariants
fun check_closing_ml_invs cid_long oid is_monitor thy =
check_invs get_closing_ml_invariants cid_long oid is_monitor thy
val check_closing_ml_invs = check_invs get_closing_ml_invariants
val ISA_prefix = "Isabelle_DOF_"
@ -1028,18 +1025,13 @@ fun ML_isa_check_term thy (term, _, pos) _ =
fun ML_isa_check_thm thy (term, _, pos) _ =
(* this works for long-names only *)
let fun check thy (name, _) = case Proof_Context.lookup_fact (Proof_Context.init_global thy) name of
NONE => err ("No Theorem:" ^name) pos
| SOME X => X
let fun check thy (name, _) = Global_Theory.check_fact thy (name, Position.none)
in ML_isa_check_generic check thy (term, pos) end
fun ML_isa_check_file thy (term, _, pos) _ =
let fun check thy (name, pos) = check_path (SOME File.check_file)
(Proof_Context.init_global thy)
(Path.current)
(name, pos);
let fun check thy (name, pos) = name |> Syntax.read_input
|> Resources.check_file (Proof_Context.init_global thy) NONE
in ML_isa_check_generic check thy (term, pos) end;
fun check_instance thy (term, _, pos) s =
@ -2015,8 +2007,8 @@ fun open_monitor_command ((((oid, pos),cid_pos), doc_attrs) : ODL_Meta_Args_Par
NONE => ISA_core.err ("You must specified a monitor class.") pos
| SOME (cid, _) => cid
val (accS, rejectS, aS) = compute_enabled_set cid thy
val info = {accepted_cids = accS, rejected_cids = rejectS, automatas = aS }
in DOF_core.add_monitor_info (Binding.make (oid, pos)) (DOF_core.Monitor_Info info) thy
in DOF_core.add_monitor_info (Binding.make (oid, pos))
(DOF_core.make_monitor_info (accS, rejectS, aS)) thy
end
in
o_m_c oid pos cid_pos doc_attrs #> create_monitor_entry oid