src/Tools/isac/BridgeJEdit/Calculation.thy
author wneuper <Walther.Neuper@jku.at>
Wed, 03 Aug 2022 18:17:27 +0200
changeset 60508 ce09935439b3
parent 60505 137227934d2e
child 60513 cecbfe45f053
permissions -rw-r--r--
cleanup
     1 (*  Title:      src/Tools/isac/BridgeJEdit/Calculation.thy
     2     Author:     Walther Neuper, JKU Linz
     3 
     4 Outer syntax for Isabelle/Isac's Calculation interactive via Isabelle/PIDE.
     5 Second trial following Makarius' definition of "problem" already existing in this repository.
     6 *)
     7 
     8 theory Calculation
     9   imports
    10   (**)"$ISABELLE_ISAC/Knowledge/Build_Thydata"   (*remove after devel.of BridgeJEdit*)
    11   (**)"$ISABELLE_ISAC/MathEngine/MathEngine"
    12   (** )"SPARK_FDL"                             ( * remove after devel.of BridgeJEdit*)
    13 keywords "Example" :: thy_decl
    14 and "Specification" "Model" "References" "Solution" 
    15 
    16 begin
    17 
    18 ML_file "parseC.sml"
    19 ML_file "preliminary.sml"
    20 
    21 section \<open>new code for Outer_Syntax Example  ...\<close>
    22 text \<open> 
    23   stepwise development starts at changeset
    24   https://hg.risc.uni-linz.ac.at/wneuper/isa/rev/9d98791b4080
    25   and can be inspected following the subsequent changesets..
    26 \<close>
    27 
    28 ML \<open>
    29 \<close> ML \<open>
    30 \<close> ML \<open>
    31 structure Data = Theory_Data
    32 (
    33   type T = Ctree.ctree option;
    34   val empty = NONE;
    35   val extend = I;
    36   fun merge _ = NONE;
    37 );
    38 
    39 val set_data = Data.put o SOME;
    40 fun the_data thy = 
    41   case Data.get thy of NONE => Ctree.EmptyPtree
    42   | SOME ctree => ctree;
    43 \<close> ML \<open>
    44 \<close> text \<open>
    45 A provisional Example collection.
    46 
    47 For notes from ISAC's discontinued Java front-end see
    48 https://isac.miraheze.org/wiki/Extend_ISAC_Knowledge#Add_an_example_in_XML_format
    49 \<close> ML \<open>
    50 \<close> ML \<open>
    51 val demo_example = ("The efficiency of an electrical coil depends on the mass of the kernel." ^
    52   "Given is a kernel with cross-sectional area determined by two rectangles" ^
    53   "of same shape as shown in the figure." ^
    54   "Given a radius r = 7, determine the length u and width v of the rectangles" ^
    55   "such that the cross-sectional area A (and thus the mass of the kernel) is a" ^
    56   "maximum. + Figure", 
    57    ([
    58    (*Problem model:*)
    59      "Constants [r = (7::real)]", "Maximum A", "AdditionalValues [u, v]",
    60      "Extremum (A = 2 * u * v - u \<up> 2)",
    61      "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]",
    62      "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]",
    63      "SideConditions [(u::real) / 2 = r * sin \<alpha>, 2 / v = r * cos \<alpha>]",
    64    (*MethodC model:*)
    65      "FunctionVariable a", "FunctionVariable b", "FunctionVariable \<alpha>",
    66      "Domain {0 <..< r}",
    67      "Domain {0 <..< r}",
    68      "Domain {0 <..< \<pi> / 2}",
    69      "ErrorBound (\<epsilon> = (0::real))"]: TermC.as_string list,
    70     ("Diff_App", ["univariate_calculus", "Optimisation"], 
    71       ["Optimisation", "by_univariate_calculus"]): References.T)): Example.T
    72 \<close> 
    73 setup \<open>KEStore_Elems.add_expls [(demo_example, ["Diff_App-No.123a"])]\<close>
    74 ML \<open>
    75 case Store.get (KEStore_Elems.get_expls @{theory}) ["Diff_App-No.123a"] ["Diff_App-No.123a"] of
    76   (_, ("Constants [r = (7::real)]" :: "Maximum A" :: "AdditionalValues [u, v]" :: _, 
    77     ("Diff_App", ["univariate_calculus", "Optimisation"] ,["Optimisation", "by_univariate_calculus"]))) 
    78     => ()
    79 | _ => error "intermed. example_store CHANGED";
    80 \<close> ML \<open>
    81 \<close> ML \<open>
    82 (*
    83   Implementation reasonable, as soo as
    84  * clarified, how "empty" items in Model are represented
    85  * parse_references_input works
    86  * switching \<Odot>\<Otimes> i_model for probl and meth works.
    87 *)
    88 fun check_input (*ctxt o-model ??*) (_: Problem.model_input) = []: I_Model.T (*TODO*)
    89 \<close> ML \<open>
    90 \<close> ML \<open>
    91 (* at the first encounter of an Example create the following data and transfer them to ctree:
    92     origin: eventually used by sub-problems, too
    93     O_Model: for check of missing items
    94     I_Model: for efficient check of user input
    95 *)
    96 \<close> ML \<open>
    97 fun init_ctree example_id =
    98     let
    99       (* TODO     vvvvvvvvvvvvvvv works only within Isac_Test_Base *)
   100       val thy = (*ThyC.get_theory thy_id*)@{theory Diff_App}
   101       val example_id' = References_Def.explode_id example_id
   102       val (_, (fmz as (model, refs as (_(*thy_id*), probl_id, _)))) =
   103         Store.get (KEStore_Elems.get_expls @{theory}) example_id' example_id' 
   104       val ctxt = ContextC.initialise' thy model;
   105 
   106       val o_model = Problem.from_store probl_id |> #ppc |> O_Model.init thy model
   107     in
   108       Ctree.Nd (Ctree.PblObj {
   109         fmz = fmz, origin = (o_model, refs, TermC.empty (*TODO handle CAS_Cmd*)),
   110         spec = References.empty, probl = I_Model.empty, meth = I_Model.empty,
   111         ctxt = ctxt, loc = (NONE, NONE), 
   112         branch = Ctree.NoBranch, ostate = Ctree.Incomplete, result = (TermC.empty, [])}, [])
   113     end
   114 \<close> ML \<open>
   115 (* 
   116 At the first call 
   117 * initialise ctree with origin from example_id
   118 * add items already input ..TODO..TODO..TODO..TODO..TODO..TODO..TODO..TODO..TODO
   119 *)
   120 \<close> ML \<open>
   121 fun update_state (example_id, (model, refs)) Ctree.EmptyPtree = 
   122     init_ctree example_id |>  update_state (example_id, (model, refs))
   123   | update_state (_, (model, (thy_id, probl_id, meth_id)))
   124     (Ctree.Nd (Ctree.PblObj {fmz, origin, spec as (othy, opbl, omet), probl = _, meth, ctxt, loc, branch, ostate, result}, 
   125       children)) =
   126       let 
   127         val _ = if thy_id <> othy then "re-parse all i_model" else "do nothing"
   128         val _ = if probl_id <> opbl then "switch from meth? check_input of probl" else "do nothing"
   129         val _ = if meth_id <> omet then "switch from probl check_input of meth" else "do nothing"
   130       in
   131         Ctree.Nd (Ctree.PblObj {
   132           fmz = fmz, origin = origin, spec = References.select_input spec (thy_id, probl_id, meth_id), 
   133           probl = check_input model, meth = meth, (*TODO +switch Problem_Ref | Method_Ref*)
   134           ctxt = ctxt, loc = loc, 
   135           branch = branch, ostate = ostate, result = result}, children)
   136       end
   137   | update_state _ _ = raise ERROR "update_state: uncovered case in fun.def"
   138 \<close> ML \<open>
   139 \<close> ML \<open>
   140 update_state: string * (Problem.model_input * References.T) -> Ctree.ctree -> Ctree.ctree
   141 \<close> ML \<open>
   142 type references_input = (string * (string * string))
   143 \<close> ML \<open>
   144 val parse_references_input =
   145   Parse.!!! ( \<^keyword>\<open>Theory_Ref\<close> -- \<^keyword>\<open>Theory_Ref\<close> |-- Parse.name --
   146   Parse.!!! ( \<^keyword>\<open>Problem_Ref\<close> -- \<^keyword>\<open>Theory_Ref\<close> |-- Parse.name --
   147   Parse.!!! ( \<^keyword>\<open>Method_Ref\<close> -- \<^keyword>\<open>Theory_Ref\<close> |-- Parse.name)))
   148 \<close> ML \<open>
   149 parse_references_input: references_input parser (*TODO: dpes not yet work,
   150   thus in Outer_Syntax.command \<^command_keyword>\<open>Example\<close> replacced by Problem.parse_cas*)
   151 \<close> ML \<open>
   152 op --  : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e;
   153 op |-- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> 'd * 'e
   154 \<close> ML \<open>
   155 \<close> ML \<open>
   156 \<close> ML \<open>
   157 fun is_empty str = 
   158   let
   159     val strl = Symbol.explode str |> rev
   160   in
   161     take (2, strl) = ["_", " "] orelse take (3, strl) = ["]", "_", "["]
   162   end
   163 \<close> ML \<open>
   164 val m2 = "0 < r";                val m2' = Symbol.explode m2 |> rev;
   165 val m3 = "Maximum _";            val m3' = Symbol.explode m3 |> rev;
   166 val m4 = "AdditionalValues [_]"; val m4' = Symbol.explode m4 |> rev;
   167 \<close> ML \<open>
   168 \<close> ML \<open>
   169 val _ =
   170   Outer_Syntax.command \<^command_keyword>\<open>Example\<close>
   171     "prepare ISAC problem type and register it to Knowledge Store"
   172     (Parse.name -- Parse.!!! ( \<^keyword>\<open>Specification\<close> -- \<^keyword>\<open>:\<close> --
   173          \<^keyword>\<open>Model\<close> -- \<^keyword>\<open>:\<close>  |-- Problem.parse_model_input --
   174          Parse.!!! ( \<^keyword>\<open>References\<close> -- \<^keyword>\<open>:\<close> |-- (*parse_references_input*) Problem.parse_cas)) >>
   175       (fn (example_id, (model_input, _(*references_input*))) =>
   176         Toplevel.theory (fn thy =>
   177           let
   178 (** )
   179 val _ = writeln ("#Example model_input: " ^ @{make_string} model_input)
   180 (*                #Example model_input:
   181     [("#Given", ["<markup>"]), ("#Where", ["<markup>"]), ("#Find", ["<markup>", "<markup>"]), ("#Relate", ["<markup>", "<markup>"])]*)
   182 val [("#Given", [m1]), ("#Where", [m2]), ("#Find", [m3, m4]), ("#Relate", [m5, m6])] = 
   183   model_input : Problem.model_input
   184 val _ = writeln ("#Example m.: " ^ m1 ^ " " ^ m2 ^ " " ^ m3 ^ " " ^ m4 ^ " " ^ m4 ^ " " ^ m5 ^ " " ^ m6)
   185 
   186 val m2' = Symbol.explode m2;
   187 val _ = writeln ("#Example Symbol.explode m2: " ^ @{make_string} m2')
   188 val _ = if (is_empty m2) then writeln "true" else writeln "false";
   189 val _ = if (is_empty m3) then writeln "true" else writeln "false";
   190 val _ = if (is_empty m4) then writeln "true" else writeln "false";
   191 (*                #Example m.: Constants [r = 7] 0 < r Maximum _ AdditionalValues [_] AdditionalValues [_] Extremum _ SideCondition [_]*)
   192 ( **)
   193             val (thy_id, (probl_id, meth_id)) = (*references_input*)("Diff_App", ("univariate_calculus/Optimisation", "Optimisation/by_univariate_calculus"))
   194             val input = update_state (example_id, (model_input, 
   195               (thy_id, References_Def.explode_id probl_id, References_Def.explode_id meth_id)))
   196                 Ctree.EmptyPtree;
   197           in set_data input thy end)));
   198 \<close> ML \<open>
   199 \<close> ML \<open>
   200 (*/----------------------------------------------------------------------------\*)
   201      Parse.name -- Parse.!!! ( \<^keyword>\<open>Specification\<close> -- \<^keyword>\<open>:\<close> --
   202          \<^keyword>\<open>Model\<close> -- \<^keyword>\<open>:\<close>  |-- Problem.parse_model_input --
   203          Parse.!!! ( \<^keyword>\<open>References\<close> -- \<^keyword>\<open>:\<close> |-- parse_references_input))
   204 (*: (string * (Problem.model_input * references_input)) parser*)
   205 (*\----------------------------------------------------------------------------/*) 
   206 : Token.T list ->
   207 (*/----------------------------------------------------------------------------\*)
   208     (string * (Problem.model_input * references_input))
   209 (*\----------------------------------------------------------------------------/*)
   210 * Token.T list
   211 (*/----------------------------------------------------------------------------\* )
   212    the type above is the type of the argument of \<open>(fn (_(*name*),..\<close> 
   213 ( *\----------------------------------------------------------------------------/*)
   214 \<close> ML \<open>
   215 \<close> ML \<open>
   216 \<close> text \<open>
   217   INVESTIGATE Problem.Outer_Syntax.command \ <^command_keyword>\<open>problem\<close>
   218   on Biegelinie.problem pbl_bieg : "Biegelinien"
   219 
   220   ML-source shows error-positions already perfectly: what types are involved?..
   221 \<close> ML \<open>
   222 Rule_Set.append_rules "empty" Rule_Set.empty [] (* ORIGINAL source = ARGUMENT IN fn *)
   223 (* (1a)
   224   Ouptut from Biegelinie.problem pbl_bieg : "Biegelinien":
   225   #problem source: Source {delimited = true, range = ({offset=38, end_offset=87, id=-1652}, {offset=87, id=-1652}), 
   226     text = "\127Rule_Set.append_rules \"empty\" Rule_Set.empty []\127"}
   227 *)
   228 \<close> ML \<open>
   229 val source = Input.empty;
   230 val ml = ML_Lex.read;
   231 \<close> ML \<open>
   232   (ml "Theory.setup (Problem.set_data (" @ ML_Lex.read_source source @ ml "))")
   233 (*= 
   234    [Text (Token (({}, {}), (Long_Ident, "Theory.setup"))), Text (Token (({}, {}), (Space, " "))), Text (Token (({}, {}), (Keyword, "("))), Text (Token (({}, {}), (Long_Ident, "Problem.set_data"))),
   235     Text (Token (({}, {}), (Space, " "))), Text (Token (({}, {}), (Keyword, "("))), Text (Token (({}, {}), (Space, " "))), Text (Token (({}, {}), (Keyword, ")"))), Text (Token (({}, {}), (Keyword, ")"))),
   236     Text (Token (({}, {}), (Space, " ")))]
   237    : ML_Lex.token Antiquote.antiquote list
   238    ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ type of (1a)
   239 *)
   240 \<close> ML \<open>
   241 ML_Context.expression (Input.pos_of source)
   242   (ml "Theory.setup (Problem.set_data (" @ ML_Lex.read_source source @ ml "))")
   243   : Context.generic -> Context.generic
   244 \<close> ML \<open>
   245 ML_Context.expression (Input.pos_of source)
   246   (ml "Theory.setup (Problem.set_data (" @ ML_Lex.read_source source @ ml "))")
   247   |> Context.theory_map
   248   : theory -> theory
   249 \<close> ML \<open>
   250 Outer_Syntax.command: Outer_Syntax.command_keyword -> string -> 
   251   (Toplevel.transition -> Toplevel.transition) parser -> unit
   252 \<close> ML \<open>
   253 \<close> text \<open>
   254   ML-source shows error-positions already perfectly: types involved are above.
   255 
   256   Input at Given does NOT show error-position, 
   257     but ISAC msg \<open>ME_Isa: thy "Isac_Knowledge" not in system\<close>
   258 
   259   Thus compare ML-source/input with model_input ...
   260 \<close> ML \<open>
   261 Problem.parse_model_input:                 Problem.model_input parser;
   262 Problem.parse_model_input: Token.src ->    Problem.model_input * Token.src;
   263 Problem.parse_model_input: Token.T list -> Problem.model_input * (Token.T list)
   264 \<close> ML \<open>
   265 []: Position.range list;
   266 []: Position.T list;
   267 \<close> ML \<open>
   268 \<close> ML \<open>
   269 (* Given: "Traegerlaenge l_l" "Streckenlast q_q" *)
   270 (* (1b)
   271   Ouptut from Biegelinie.problem pbl_bieg : "Biegelinien":
   272   #problem model_input: [("#Given", ["<markup>", "<markup>"]), ("#Find", ["<markup>"]), ("#Relate", ["<markup>"])] 
   273                         : Problem.model_input
   274   type of  (1b) --------^^^^^^^^^^^^^^^^^^^^^^> contains ? Markup.T ?..
   275 
   276   #problem m1: Traegerlaenge l_l
   277   #problem m2: Streckenlast q_q 
   278   I this -----^^^______________^^^ Markup.T ?
   279 *)
   280 \<close> ML \<open> (** investigate parse_model_input **)
   281 \<close> ML \<open> (* how get a Token.T list without Outer_Syntax.command ? *)
   282 \<close> ML \<open>
   283 Parse.term:                 string parser;
   284 Parse.term: Token.T list -> string * Token.T list;
   285 \<close> ML \<open>
   286 \<close> ML \<open> (* relation Token.T -- Markup.T ??? *)
   287 \<close> ML \<open>
   288 val ctxt = Proof_Context.init_global @{theory};
   289 val given_1 = Syntax.read_term ctxt "Traegerlaenge l_l" (* as in parse_tagged_terms *)
   290 (*  ^^^^^^^ : term  -- why then evaluates parse_model_input to markup ? *)
   291 \<close> ML \<open>
   292 \<close> ML \<open>
   293 Markup.literal_fact "Traegerlaenge l_l" = ("literal_fact", [("name", "Traegerlaenge l_l")])
   294 \<close> ML \<open>
   295 \<close> ML \<open>
   296 \<close> ML \<open>
   297 \<close>
   298 
   299 
   300 end