src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Sat, 03 Apr 2021 15:27:52 +0200
changeset 60181 2654f8196c36
parent 60149 f01072d28542
child 60192 4c7c15750166
permissions -rw-r--r--
review and update directories in Build_Isac
walther@59632
     1
(*  Title:  build the Isac-Kernel: MathEngine & Knowledge
neuper@38057
     2
    Author: Walther Neuper, TU Graz, 100808
neuper@38051
     3
   (c) due to copyright terms
neuper@38051
     4
neuper@52062
     5
For creating a heap image of isac see ~~/ROOT.
walther@59632
     6
For debugging see text at begin below, e.g. theory dependencies:
walther@59632
     7
  ~$ evince file:///home/wneuper/.isabelle/isabisac/browser_info/Unsorted/Isac/session_graph.pdf &
neuper@52062
     8
wneuper@59586
     9
ATTENTION: no errors in this theory do not mean that there are no errors in Isac ..
walther@59866
    10
.. open theories collecting files from folders: BaseDefinitions.thy, ProgLang.thy etc.
walther@59632
    11
Errors are rigorously detected by isabelle build.
neuper@37906
    12
*)
neuper@37906
    13
walther@59998
    14
theory Build_Isac
wneuper@59206
    15
imports
walther@59887
    16
(*  theory Know_Store imports Complex_Main
walther@60181
    17
    ~~/src/Tools/isac/BaseDefinitions
walther@59632
    18
      ML_file libraryC.sml
walther@59865
    19
      ML_file theoryC.sml
walther@59963
    20
      ML_file unparseC.sml
walther@59850
    21
      ML_file "rule-def.sml"
walther@59963
    22
      ML_file "thmC-def.sml"
walther@59919
    23
      ML_file "eval-def.sml"
walther@59858
    24
      ML_file "rewrite-order.sml"
walther@59632
    25
      ML_file rule.sml
walther@59963
    26
      ML_file "error-pattern-def.sml"
walther@59852
    27
      ML_file "rule-set.sml"
walther@59963
    28
      
walther@59963
    29
      ML_file "store.sml"
walther@59963
    30
      ML_file "check-unique.sml"
walther@59963
    31
      ML_file "specification.sml"
walther@59963
    32
      ML_file "model-pattern.sml"
walther@59963
    33
      ML_file "problem-def.sml"
walther@59963
    34
      ML_file "method-def.sml"
walther@59963
    35
      ML_file "cas-def.sml"
walther@59963
    36
      ML_file "thy-write.sml"
walther@59887
    37
  theory BaseDefinitions imports Know_Store
walther@60181
    38
  ~~/src/Tools/isac/BaseDefinitions
walther@59632
    39
    ML_file termC.sml
walther@59963
    40
    ML_file substitution.sml
walther@59632
    41
    ML_file contextC.sml
walther@59687
    42
    ML_file environment.sml
walther@60077
    43
( ** )    "BaseDefinitions/BaseDefinitions"( **)
wneuper@59586
    44
walther@59866
    45
(*    theory Calculate imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
walther@60181
    46
      ~~/src/Tools/isac/ProgLang
walther@59902
    47
        ML_file evaluate.sml
walther@59773
    48
walther@59866
    49
      theory ListC imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
walther@60181
    50
      ~~/src/Tools/isac/ProgLang
walther@59632
    51
    theory Prog_Expr imports Calculate ListC
walther@60181
    52
    ~~/src/Tools/isac/ProgLang
walther@59866
    53
      theory Program imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
walther@59866
    54
      theory Prog_Tac imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
walther@59866
    55
      theory Tactical imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
walther@59632
    56
    theory Auto_Prog imports Program Prog_Tac Tactical begin
walther@60181
    57
    ~~/src/Tools/isac/ProgLang
walther@59632
    58
  theory ProgLang imports Prog_Expr Auto_Prog
walther@60181
    59
  ~~/src/Tools/isac/ProgLang
walther@60077
    60
( ** )    "ProgLang/ProgLang"( **)
wneuper@59600
    61
(*
walther@59674
    62
  theory MathEngBasic imports
walther@59674
    63
    "~~/src/Tools/isac/ProgLang/ProgLang" "~~/src/Tools/isac/Specify/Input_Descript"
walther@60181
    64
  ~~/src/Tools/isac/MathEngBasic
walther@59875
    65
    ML_file thmC.sml
walther@59963
    66
    ML_file problem.sml
walther@59963
    67
    ML_file method.sml
walther@59963
    68
    ML_file "cas-command.sml"
walther@59963
    69
  
walther@59865
    70
    ML_file rewrite.sml
walther@59963
    71
  
walther@59963
    72
    ML_file "model-def.sml"
walther@59963
    73
    ML_file "istate-def.sml"
walther@59674
    74
    ML_file "calc-tree-elem.sml"
walther@59963
    75
    ML_file "pre-conds-def.sml"
walther@59963
    76
  
walther@59632
    77
    ML_file tactic.sml
walther@59963
    78
    ML_file applicable.sml
walther@59963
    79
  
walther@59777
    80
    ML_file position.sml
walther@59674
    81
    ML_file "ctree-basic.sml"
walther@59674
    82
    ML_file "ctree-access.sml"
walther@59674
    83
    ML_file "ctree-navi.sml"
walther@59674
    84
    ML_file ctree.sml
walther@59963
    85
  
walther@59963
    86
    ML_file "state-steps.sml"
walther@59777
    87
    ML_file calculation.sml
walther@60077
    88
( ** )    "MathEngBasic/MathEngBasic"( **)
walther@59674
    89
(*
walther@59866
    90
    theory Input_Descript imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
walther@60181
    91
    ~~/src/Tools/isac/Specify
walther@59674
    92
  theory Specify imports "~~/src/Tools/isac/ProgLang/ProgLang" Input_Descript
walther@60181
    93
  ~~/src/Tools/isac/Specify
walther@59963
    94
    ML_file formalise.sml
walther@59963
    95
    ML_file "o-model.sml"
walther@59963
    96
    ML_file "i-model.sml"
walther@59963
    97
    ML_file "p-model.sml"
walther@59963
    98
    ML_file model.sml
walther@59963
    99
    ML_file "pre-conditions.sml"
walther@59963
   100
    ML_file refine.sml
walther@59963
   101
    ML_file "mstools.sml" (*..TODO review*)
walther@59963
   102
    ML_file ptyps.sml     (*..TODO review*)
walther@59963
   103
    ML_file "test-out.sml"
walther@59963
   104
    ML_file "specify-step.sml"
walther@59963
   105
    ML_file calchead.sml  (*..TODO review*)
walther@59963
   106
    ML_file "input-calchead.sml"
walther@59777
   107
    ML_file "step-specify.sml"
walther@59777
   108
    ML_file specify.sml
walther@60077
   109
( ** )    "Specify/Specify"( **)
wneuper@59600
   110
(*
walther@59632
   111
  theory Interpret imports "~~/src/Tools/isac/Specify/Specify"
walther@60181
   112
  ~~/src/Tools/isac/Interpret
walther@59777
   113
    ML_file istate.sml
walther@59963
   114
    ML_file "sub-problem.sml"
walther@59963
   115
    ML_file "thy-read.sml"
walther@59963
   116
    ML_file "solve-step.sml"
walther@59963
   117
    ML_file "error-pattern.sml"
walther@59963
   118
    ML_file derive.sml
walther@59794
   119
    ML_file "li-tool.sml"
walther@59632
   120
    ML_file "lucas-interpreter.sml"
walther@59794
   121
    ML_file "step-solve.sml"
walther@60077
   122
( ** )    "Interpret/Interpret"( **)
wneuper@59600
   123
(*
walther@60077
   124
  theory MathEngine imports Interpret.Interpret
walther@60181
   125
  ~~/src/Tools/isac/MathEngine
walther@59833
   126
    ML_file "fetch-tactics.sml"
walther@59632
   127
    ML_file solve.sml
walther@59794
   128
    ML_file step.sml
walther@59833
   129
    ML_file "detail-step.sml"
walther@59634
   130
    ML_file "mathengine-stateless.sml"
walther@59632
   131
    ML_file messages.sml
walther@59632
   132
    ML_file states.sml
walther@60077
   133
( ** )    "MathEngine/MathEngine"( **)
wneuper@59600
   134
(*
walther@59814
   135
  theory Test_Code imports "~~/src/Tools/isac/MathEngine/MathEngine"
walther@60181
   136
  ~~/src/Tools/isac/Test_Code
walther@59814
   137
    ML_file "test-code.sml"
walther@60077
   138
( ** )    "Test_Code/Test_Code"( **)
walther@59814
   139
(*
walther@59632
   140
  theory BridgeLibisabelle imports "~~/src/Tools/isac/MathEngine/MathEngine"
walther@60181
   141
  ~~/src/Tools/isac/BridgeLibisabelle
walther@59963
   142
    ML_file "thy-present.sml"
walther@59963
   143
    ML_file mathml.sml   
walther@59632
   144
    ML_file datatypes.sml
walther@59632
   145
    ML_file "pbl-met-hierarchy.sml"
walther@59632
   146
    ML_file "thy-hierarchy.sml" 
walther@59632
   147
    ML_file "interface-xml.sml"
walther@59632
   148
    ML_file interface.sml
walther@60077
   149
( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
walther@60044
   150
(*
walther@60181
   151
    theory Calculation imports Calculation (*preliminary for starting Isabelle/Isac*)
walther@60181
   152
    ~~/src/Tools/isac/BridgeJEdit
walther@60181
   153
      ML_file parseC.sml
walther@60181
   154
      ML_file preliminary.sml
walther@60181
   155
  theory BridgeJEdit imports Calculation (*preliminary for starting Isabelle/Isac*)
walther@60181
   156
  ~~/src/Tools/isac/BridgeJEdit
walther@60181
   157
( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
walther@60181
   158
(*
walther@60044
   159
    theory Isac imports  "~~/src/Tools/isac/MathEngine/MathEngine"
walther@60044
   160
      ML_file parseC.sml
walther@60044
   161
  theory BridgeJEdit imports Isac    
walther@60149
   162
( **)    "BridgeJEdit/BridgeJEdit"  (*DEactivate after devel.of BridgeJEdit*)
wneuper@59147
   163
walther@59612
   164
          "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
wneuper@59469
   165
wneuper@59601
   166
(*//-----------------------------------------------------------------------------------------\\*)
wneuper@59601
   167
(*\\-----------------------------------------------------------------------------------------//*)
neuper@48763
   168
begin
neuper@48760
   169
wneuper@59472
   170
text \<open>
neuper@55276
   171
  show theory dependencies using the graph browser, 
neuper@55276
   172
  open "browser_info/HOL/Isac/session.graph"
neuper@55276
   173
  and proceed from the ancestors towards the siblings.
wneuper@59472
   174
\<close>
neuper@55276
   175
wneuper@59472
   176
section \<open>check presence of definitions from directories\<close>
neuper@48763
   177
wneuper@59263
   178
(*declare [[ML_print_depth = 999]]*)
wneuper@59472
   179
ML \<open>
wneuper@59472
   180
\<close> ML \<open>
wneuper@59587
   181
\<close> ML \<open>
wneuper@59472
   182
\<close>
walther@59902
   183
ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
wneuper@59472
   184
ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
walther@59953
   185
ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
walther@59814
   186
ML \<open>Test_Code.me;\<close>
walther@59613
   187
text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
walther@59801
   188
ML \<open>prog_expr\<close>
wneuper@59562
   189
                                                         
walther@59603
   190
ML \<open>Prog_Expr.eval_occurs_in\<close>
walther@59603
   191
ML \<open>@{thm last_thmI}\<close>
walther@60077
   192
(** )ML \<open>@{thm Querkraft_Belastung}\<close>( *exception FAIL NONE raised (line 161 of "General/scan.ML")*)
neuper@41905
   193
walther@59904
   194
ML \<open>Check_Unique.on := false;\<close>
wneuper@59472
   195
ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
wneuper@59592
   196
ML \<open>@{theory "Isac_Knowledge"}\<close>
wneuper@59472
   197
ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
walther@59997
   198
  ERROR: app_py: not found: ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]*)\<close>
neuper@42412
   199
wneuper@59472
   200
section \<open>State of approaching Isabelle by Isac\<close>
wneuper@59472
   201
text \<open>
neuper@55431
   202
  Mathias Lehnfeld gives the following list in his thesis in section 
neuper@55431
   203
  4.2.3 Relation to Ongoing Isabelle Development.
wneuper@59472
   204
\<close>
wneuper@59472
   205
subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
wneuper@59472
   206
text \<open>
walther@59887
   207
  REPLACE BY Know_Store... (has been overlooked)
neuper@55433
   208
    calcelems.sml:val rew_ord' = Unsynchronized.ref ...
neuper@55433
   209
  KEEP FOR TRACING
walther@59901
   210
    rewrite.sml:val Rewrite.trace_on = Unsynchronized.ref false;
walther@59901
   211
    rewrite.sml:val depth = Unsynchronized.ref 99999;
walther@59901
   212
    rewrite.sml:val lim_rewrite = Unsynchronized.ref 99999;
walther@59901
   213
    rewrite.sml:val lim_deriv = Unsynchronized.ref 100;
walther@59901
   214
    Interpret/rewtools.sml:val LItool.trace = Unsynchronized.ref false;
neuper@55433
   215
  KEEP FOR EASIER DEVELOPMENT
neuper@55433
   216
    calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
neuper@55433
   217
  KEEP FOR DEMOS
neuper@55433
   218
    Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
neuper@55433
   219
    Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
neuper@55433
   220
    Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
wneuper@59472
   221
\<close>
wneuper@59472
   222
subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
wneuper@59472
   223
subsection \<open>(2) Make Isac’s programming language usable\<close>
wneuper@59472
   224
subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
wneuper@59472
   225
text \<open>
neuper@55431
   226
  In 2002 isac already strived for floating point numbers. Since that time
neuper@55431
   227
  isac represents numerals as "Free", see below (*1*). These experiments are
neuper@55431
   228
  unsatisfactory with respect to logical soundness.
neuper@55431
   229
  Since Isabelle now has started to care about floating point numbers, it is high 
neuper@55431
   230
  time to adopt these together with the other numerals. Isabelle2012/13's numerals
neuper@55431
   231
  are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
neuper@55431
   232
  
neuper@55431
   233
  The transition from "Free" to standard numerals is a task to be scheduled for 
neuper@55431
   234
  several weeks. The urgency of this task follows from the experience, 
neuper@55431
   235
  that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
neuper@55431
   236
  some of the long identifiers of theorems which would tremendously simplify
neuper@55431
   237
  building a hierarchy of theorems according to (1.2), see (*2*) below.
wneuper@59472
   238
\<close>
wneuper@59472
   239
ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
neuper@55484
   240
wneuper@59472
   241
subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
wneuper@59472
   242
subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
neuper@55431
   243
neuper@37906
   244
end