src/Tools/isac/BridgeJEdit/Calculation.thy
author wneuper <Walther.Neuper@jku.at>
Mon, 20 Jun 2022 11:51:12 +0200
changeset 60467 d0b31d83cfad
parent 60458 af7735fd252f
child 60469 89e1d8a633bb
permissions -rw-r--r--
initialise state from intermed. example_store
walther@60096
     1
(*  Title:      src/Tools/isac/BridgeJEdit/Calculation.thy
walther@60095
     2
    Author:     Walther Neuper, JKU Linz
walther@60095
     3
Walther@60431
     4
Outer syntax for Isabelle/Isac's Calculation interactive via Isabelle/PIDE.
Walther@60431
     5
Second trial following Makarius' definition of "problem" already existing in this repository.
walther@60095
     6
*)
walther@60095
     7
walther@60095
     8
theory Calculation
Walther@60434
     9
  imports
Walther@60434
    10
  (**)"$ISABELLE_ISAC/Knowledge/Build_Thydata"   (*remove after devel.of BridgeJEdit*)
Walther@60434
    11
  (**)"$ISABELLE_ISAC/MathEngine/MathEngine"
Walther@60434
    12
  (** )"SPARK_FDL"                             ( * remove after devel.of BridgeJEdit*)
Walther@60450
    13
keywords "Example" :: thy_decl
Walther@60451
    14
and "Problem" "Specification" "Model" "References" "Solution" 
Walther@60431
    15
walther@60095
    16
begin
wenzelm@60200
    17
Walther@60431
    18
ML_file "parseC.sml"
Walther@60431
    19
ML_file "preliminary.sml"
walther@60095
    20
Walther@60434
    21
section \<open>new code for Outer_Syntax Example  ...\<close>
Walther@60441
    22
text \<open> 
Walther@60456
    23
  stepwise development starts at changeset
Walther@60441
    24
  https://hg.risc.uni-linz.ac.at/wneuper/isa/rev/9d98791b4080
Walther@60456
    25
  and can be inspected following the subsequent changesets..
Walther@60434
    26
\<close>
walther@60123
    27
Walther@60434
    28
ML \<open>
Walther@60434
    29
\<close> ML \<open>
Walther@60434
    30
\<close> ML \<open>
Walther@60434
    31
structure Data = Theory_Data
Walther@60434
    32
(
Walther@60442
    33
  type T = Ctree.ctree option;
Walther@60434
    34
  val empty = NONE;
Walther@60434
    35
  val extend = I;
Walther@60434
    36
  fun merge _ = NONE;
Walther@60434
    37
);
walther@60150
    38
Walther@60437
    39
val set_data = Data.put o SOME;
Walther@60457
    40
fun the_data thy = 
Walther@60457
    41
  case Data.get thy of NONE => Ctree.EmptyPtree
Walther@60457
    42
  | SOME ctree => ctree;
Walther@60457
    43
\<close> ML \<open>
Walther@60457
    44
\<close> text \<open>
Walther@60457
    45
A provisional Example collection.
Walther@60434
    46
Walther@60457
    47
For notes from ISAC's discontinued Java front-end see
Walther@60457
    48
https://isac.miraheze.org/wiki/Extend_ISAC_Knowledge#Add_an_example_in_XML_format
Walther@60434
    49
\<close> ML \<open>
Walther@60457
    50
type description = string
Walther@60457
    51
type example = (description * Formalise.T)
Walther@60467
    52
type example_id = string
Walther@60467
    53
val example_store = Unsynchronized.ref ([]: (example_id * example) list)
Walther@60467
    54
fun get_example example_id =
Walther@60467
    55
  case AList.lookup op= (!example_store) example_id of
Walther@60467
    56
    NONE => raise ERROR ("Examnple " ^ quote example_id ^ " not found")
Walther@60467
    57
  | SOME expl => expl
Walther@60434
    58
\<close> ML \<open>
Walther@60467
    59
get_example: example_id -> example;
Walther@60467
    60
\<close> ML \<open>
Walther@60467
    61
example_store := AList.update op= ("Diff_App/No.123 a", 
Walther@60457
    62
 ("The efficiency of an electrical coil depends on the mass of the kernel." ^
Walther@60457
    63
  "Given is a kernel with cross-sectional area determined by two rectangles" ^
Walther@60457
    64
  "of same shape as shown in the figure." ^
Walther@60457
    65
  "Given a radius r = 7, determine the length u and width v of the rectangles" ^
Walther@60457
    66
  "such that the cross-sectional area A (and thus the mass of the kernel) is a" ^
Walther@60457
    67
  "maximum. + Figure", 
Walther@60467
    68
   ([
Walther@60467
    69
   (*Problem model:*)
Walther@60467
    70
     "Constants [r = (7::real)]", "Maximum A", "AdditionalValues [u, v]",
Walther@60467
    71
     "Extremum (A = 2 * u * v - u \<up> 2)",
Walther@60467
    72
     "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]",
Walther@60467
    73
     "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]",
Walther@60467
    74
     "SideConditions [(u::real) / 2 = r * sin \<alpha>, 2 / v = r * cos \<alpha>]",
Walther@60467
    75
   (*MethodC model:*)
Walther@60467
    76
     "FunctionVariable a", "FunctionVariable b", "FunctionVariable \<alpha>",
Walther@60467
    77
     "Domain {0 <..< r}",
Walther@60467
    78
     "Domain {0 <..< r}",
Walther@60467
    79
     "Domain {0 <..< \<pi> / 2}",
Walther@60467
    80
     "ErrorBound (\<epsilon> = (0::real))"]: TermC.as_string list,
Walther@60457
    81
    ("Diff_App", ["univariate_calculus", "Optimisation"], 
Walther@60467
    82
      ["Optimisation", "by_univariate_calculus"]): References.T)): example) (!example_store)
Walther@60457
    83
\<close> ML \<open>
Walther@60467
    84
case get_example "Diff_App/No.123 a" of
Walther@60467
    85
  (_, ("Constants [r = (7::real)]" :: "Maximum A" :: "AdditionalValues [u, v]" :: _, 
Walther@60457
    86
    ("Diff_App", ["univariate_calculus", "Optimisation"] ,["Optimisation", "by_univariate_calculus"]))) 
Walther@60457
    87
    => ()
Walther@60457
    88
| _ => error "intermed. example_store CHANGED";
Walther@60457
    89
\<close> ML \<open>
Walther@60434
    90
\<close> ML \<open>
Walther@60454
    91
fun check_input (*ctxt o-model ??*) (_: Problem.model_input) = []: I_Model.T
Walther@60434
    92
\<close> ML \<open>
Walther@60467
    93
(*at the first encounter of an Example
Walther@60467
    94
  transfer Example to ctree:
Walther@60467
    95
    origin: eventually used by sub-problems, too
Walther@60467
    96
    O_Model: for check of missing items
Walther@60467
    97
    I_Model: for efficient check of user input
Walther@60467
    98
*)
Walther@60467
    99
\<close> ML \<open>
Walther@60467
   100
\<close> ML \<open>
Walther@60467
   101
\<close> ML \<open>
Walther@60467
   102
\<close> ML \<open>
Walther@60467
   103
\<close> ML \<open>
Walther@60467
   104
fun init_ctree example_id =
Walther@60467
   105
    let
Walther@60467
   106
      val (_, (fmz as (model, refs as (_(*thy_id*), probl_id, _)))) = get_example example_id
Walther@60467
   107
      val o_model =    (* TODO works only within Isac_Test_Base ---vvvvvvvvvvvvvvv *)
Walther@60467
   108
      (*Problem.from_store probl_id |> #ppc |> O_Model.init model (ThyC.get_theory thy_id)*)
Walther@60467
   109
        Problem.from_store probl_id |> #ppc |> O_Model.init model @{theory Diff_App}
Walther@60467
   110
    in
Walther@60467
   111
      Ctree.Nd (Ctree.PblObj {
Walther@60467
   112
        fmz = fmz, origin = (o_model, refs, TermC.empty (*TODO handle CAS_Cmd*)),
Walther@60467
   113
        spec = References.empty, probl = I_Model.empty, meth = I_Model.empty,
Walther@60467
   114
        ctxt = ContextC.empty, loc = (NONE, NONE), 
Walther@60467
   115
        branch = Ctree.NoBranch, ostate = Ctree.Incomplete, result = (TermC.empty, [])}, [])
Walther@60467
   116
    end
Walther@60467
   117
\<close> ML \<open>
Walther@60467
   118
\<close> ML \<open>
Walther@60467
   119
fun update_state (example_id, _) Ctree.EmptyPtree = init_ctree example_id
Walther@60467
   120
  (*TODO +switch RProblem | RMethod*)
Walther@60467
   121
  | update_state (_, (model, (thy_id, probl_id, meth_id)))
Walther@60454
   122
    (Ctree.Nd (Ctree.PblObj {fmz, origin, spec = _, probl = _, meth, ctxt, loc, branch, ostate, result}, 
Walther@60443
   123
      children)) =
Walther@60443
   124
    Ctree.Nd (Ctree.PblObj {
Walther@60443
   125
      fmz = fmz, origin = origin,
Walther@60454
   126
      spec = (thy_id, probl_id, meth_id), probl = check_input model, meth = meth,
Walther@60443
   127
      ctxt = ctxt, loc = loc, 
Walther@60443
   128
      branch = branch, ostate = ostate, result = result}, children)
Walther@60443
   129
  | update_state _ _ = raise ERROR "update_state: uncovered case in fun.def."
Walther@60443
   130
\<close> ML \<open>
Walther@60467
   131
update_state: example_id * (Problem.model_input * References.T) -> Ctree.ctree -> Ctree.ctree
Walther@60443
   132
\<close> ML \<open>
Walther@60455
   133
type references_input = (string * (string * string))
Walther@60455
   134
\<close> ML \<open>
Walther@60455
   135
val parse_references_input =
Walther@60455
   136
  Parse.!!! ( \<^keyword>\<open>Theory_Ref\<close> -- \<^keyword>\<open>Theory_Ref\<close> |-- Parse.name --
Walther@60455
   137
  Parse.!!! ( \<^keyword>\<open>Problem_Ref\<close> -- \<^keyword>\<open>Theory_Ref\<close> |-- Parse.name --
Walther@60455
   138
  Parse.!!! ( \<^keyword>\<open>Method_Ref\<close> -- \<^keyword>\<open>Theory_Ref\<close> |-- Parse.name)))
Walther@60455
   139
\<close> ML \<open>
Walther@60457
   140
parse_references_input: references_input parser (*TODO: dpes not yet work,
Walther@60457
   141
  thus in Outer_Syntax.command \<^command_keyword>\<open>Example\<close> replacced by Problem.parse_cas*)
Walther@60455
   142
\<close> ML \<open>
Walther@60455
   143
op --  : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e;
Walther@60455
   144
op |-- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> 'd * 'e
Walther@60455
   145
\<close> ML \<open>
Walther@60455
   146
\<close> ML \<open>
Walther@60443
   147
\<close> ML \<open>
Walther@60434
   148
val _ =
Walther@60434
   149
  Outer_Syntax.command \<^command_keyword>\<open>Example\<close>
Walther@60434
   150
    "prepare ISAC problem type and register it to Knowledge Store"
Walther@60452
   151
    (Parse.name -- Parse.!!! (  \<^keyword>\<open>Problem\<close> |-- Parse.name -- 
Walther@60455
   152
       Parse.!!! ( \<^keyword>\<open>Specification\<close> -- \<^keyword>\<open>:\<close> --
Walther@60455
   153
         \<^keyword>\<open>Model\<close> -- \<^keyword>\<open>:\<close>  |-- Problem.parse_model_input --
Walther@60455
   154
         Parse.!!! ( \<^keyword>\<open>References\<close> -- \<^keyword>\<open>:\<close> |-- (*parse_references_input*)Problem.parse_cas))) >>
Walther@60467
   155
      (fn (example_id, (_(*probl_id'*), (model_input, _(*references_input*)))) =>
Walther@60434
   156
        Toplevel.theory (fn thy =>
Walther@60434
   157
          let
Walther@60455
   158
            val (thy_id, (probl_id, meth_id)) = (*references_input*)("Diff_App", ("univariate_calculus/Optimisation", "Optimisation/by_univariate_calculus"))
Walther@60467
   159
            val input = update_state (example_id, (model_input, 
Walther@60467
   160
              (thy_id, References_Def.explode_id probl_id, References_Def.explode_id meth_id)))
Walther@60455
   161
                Ctree.EmptyPtree;
Walther@60443
   162
          in set_data input thy end)));
Walther@60434
   163
\<close> ML \<open>
Walther@60447
   164
\<close> ML \<open>
Walther@60445
   165
(*/----------------------------------------------------------------------------\*)
Walther@60452
   166
     Parse.name -- Parse.!!! (  \<^keyword>\<open>Problem\<close> |-- Parse.name -- 
Walther@60455
   167
       Parse.!!! ( \<^keyword>\<open>Specification\<close> -- \<^keyword>\<open>:\<close> --
Walther@60455
   168
         \<^keyword>\<open>Model\<close> -- \<^keyword>\<open>:\<close>  |-- Problem.parse_model_input --
Walther@60455
   169
         Parse.!!! ( \<^keyword>\<open>References\<close> -- \<^keyword>\<open>:\<close> |-- parse_references_input)))
Walther@60455
   170
: (string * (string * (Problem.model_input * references_input))) parser
Walther@60455
   171
(*\----------------------------------------------------------------------------/*) 
Walther@60445
   172
: Token.T list ->
Walther@60445
   173
(*/----------------------------------------------------------------------------\*)
Walther@60455
   174
    (string * (string * (Problem.model_input * references_input)))
Walther@60445
   175
(*\----------------------------------------------------------------------------/*)
Walther@60445
   176
* Token.T list
Walther@60445
   177
(*/----------------------------------------------------------------------------\* )
Walther@60445
   178
   the type above is the type of the argument of \<open>(fn (_(*name*),..\<close> 
Walther@60445
   179
( *\----------------------------------------------------------------------------/*)
Walther@60434
   180
\<close> ML \<open>
Walther@60434
   181
\<close> ML \<open>
Walther@60434
   182
\<close> 
Walther@60434
   183
Walther@60434
   184
Walther@60434
   185
end