src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Mon, 24 Feb 2020 17:51:26 +0100
changeset 59814 665dd868d4e2
parent 59801 17d807bf28fb
child 59833 9331e61f55dd
permissions -rw-r--r--
prep.: add test-code and test, cleanup
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 ..
wneuper@59586
    10
.. open theories collecting files from folders: CalcElements.thy, ProgLang.thy etc.
walther@59632
    11
Errors are rigorously detected by isabelle build.
neuper@37906
    12
*)
neuper@37906
    13
wneuper@59347
    14
theory Build_Isac 
wneuper@59206
    15
imports
walther@59632
    16
(*  theory KEStore imports Complex_Main
walther@59632
    17
      ML_file libraryC.sml
walther@59632
    18
      ML_file rule.sml
walther@59632
    19
      ML_file calcelems.sml
walther@59632
    20
  theory CalcElements imports KEStore
walther@59632
    21
    ML_file termC.sml
walther@59632
    22
    ML_file contextC.sml
walther@59687
    23
    ML_file environment.sml
wneuper@59586
    24
*)        "CalcElements/CalcElements"
wneuper@59586
    25
walther@59632
    26
(*    theory Calculate imports "~~/src/Tools/isac/CalcElements/CalcElements"
walther@59632
    27
        ML_file calculate.sml
walther@59773
    28
walther@59632
    29
      theory ListC imports "~~/src/Tools/isac/CalcElements/CalcElements"
walther@59632
    30
    theory Prog_Expr imports Calculate ListC
walther@59632
    31
      theory Program imports "~~/src/Tools/isac/CalcElements/CalcElements"
walther@59632
    32
      theory Prog_Tac imports "~~/src/Tools/isac/CalcElements/CalcElements"
walther@59632
    33
      theory Tactical imports "~~/src/Tools/isac/CalcElements/CalcElements"
walther@59632
    34
    theory Auto_Prog imports Program Prog_Tac Tactical begin
walther@59632
    35
  theory ProgLang imports Prog_Expr Auto_Prog
walther@59632
    36
    ML_file rewrite.sml
neuper@52062
    37
*)        "ProgLang/ProgLang"
wneuper@59600
    38
(*
walther@59674
    39
  theory MathEngBasic imports
walther@59674
    40
    "~~/src/Tools/isac/ProgLang/ProgLang" "~~/src/Tools/isac/Specify/Input_Descript"
walther@59674
    41
    ML_file "calc-tree-elem.sml"
walther@59632
    42
    ML_file model.sml
walther@59632
    43
    ML_file mstools.sml
walther@59632
    44
    ML_file "specification-elems.sml"
walther@59777
    45
    ML_file "istate-def.sml"
walther@59632
    46
    ML_file tactic.sml
walther@59777
    47
    ML_file position.sml
walther@59674
    48
    ML_file "ctree-basic.sml"
walther@59674
    49
    ML_file "ctree-access.sml"
walther@59674
    50
    ML_file "ctree-navi.sml"
walther@59674
    51
    ML_file ctree.sml
walther@59779
    52
  (*ML_file tactic.sml*)
walther@59777
    53
    ML_file calculation.sml
walther@59674
    54
*)        "MathEngBasic/MathEngBasic"
walther@59674
    55
(*
walther@59674
    56
    theory Input_Descript imports "~~/src/Tools/isac/CalcElements/CalcElements"
walther@59674
    57
  theory Specify imports "~~/src/Tools/isac/ProgLang/ProgLang" Input_Descript
walther@59632
    58
    ML_file ptyps.sml
walther@59632
    59
    ML_file generate.sml
walther@59632
    60
    ML_file calchead.sml
walther@59632
    61
    ML_file appl.sml
walther@59777
    62
    ML_file "step-specify.sml"
walther@59777
    63
    ML_file specify.sml
wneuper@59595
    64
*)        "Specify/Specify"
wneuper@59600
    65
(*
walther@59632
    66
  theory Interpret imports "~~/src/Tools/isac/Specify/Specify"
walther@59777
    67
    ML_file istate.sml
walther@59632
    68
    ML_file rewtools.sml
walther@59794
    69
    ML_file "li-tool.sml"
walther@59632
    70
    ML_file inform.sml
walther@59632
    71
    ML_file "lucas-interpreter.sml"
walther@59794
    72
    ML_file "step-solve.sml"
wneuper@59600
    73
*)        "Interpret/Interpret"
wneuper@59600
    74
(*
walther@59632
    75
  theory MathEngine imports "~~/src/Tools/isac/Interpret/Interpret"
walther@59632
    76
    ML_file solve.sml
walther@59794
    77
    ML_file step.sml
walther@59634
    78
    ML_file "mathengine-stateless.sml"
walther@59632
    79
    ML_file messages.sml
walther@59632
    80
    ML_file states.sml
wneuper@59600
    81
*)        "MathEngine/MathEngine"
wneuper@59600
    82
(*
walther@59814
    83
  theory Test_Code imports "~~/src/Tools/isac/MathEngine/MathEngine"
walther@59814
    84
    ML_file "test-code.sml"
walther@59814
    85
*)        "Test_Code/Test_Code"
walther@59814
    86
(*
walther@59632
    87
  theory BridgeLibisabelle imports "~~/src/Tools/isac/MathEngine/MathEngine"
walther@59632
    88
    ML_file mathml.sml
walther@59632
    89
    ML_file datatypes.sml
walther@59632
    90
    ML_file "pbl-met-hierarchy.sml"
walther@59632
    91
    ML_file "thy-hierarchy.sml" 
walther@59632
    92
    ML_file "interface-xml.sml"
walther@59632
    93
    ML_file interface.sml
walther@59614
    94
*)        "BridgeLibisabelle/BridgeLibisabelle"
wneuper@59147
    95
walther@59612
    96
          "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
wneuper@59469
    97
wneuper@59601
    98
(*//-----------------------------------------------------------------------------------------\\*)
wneuper@59601
    99
(*\\-----------------------------------------------------------------------------------------//*)
neuper@48763
   100
begin
neuper@48760
   101
wneuper@59472
   102
text \<open>
neuper@55276
   103
  show theory dependencies using the graph browser, 
neuper@55276
   104
  open "browser_info/HOL/Isac/session.graph"
neuper@55276
   105
  and proceed from the ancestors towards the siblings.
wneuper@59472
   106
\<close>
neuper@55276
   107
wneuper@59472
   108
section \<open>check presence of definitions from directories\<close>
neuper@48763
   109
wneuper@59263
   110
(*declare [[ML_print_depth = 999]]*)
wneuper@59472
   111
ML \<open>
wneuper@59472
   112
\<close> ML \<open>
walther@59637
   113
parseNEW @{context} "xxx #> aaa bbb"
walther@59637
   114
\<close> ML \<open>
walther@59637
   115
#> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c
wneuper@59587
   116
\<close> ML \<open>
wneuper@59472
   117
\<close>
walther@59773
   118
ML \<open>Num_Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
wneuper@59472
   119
ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
wneuper@59595
   120
ML \<open>Input_Descript.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
walther@59814
   121
ML \<open>Test_Code.me;\<close>
walther@59613
   122
text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
walther@59801
   123
ML \<open>prog_expr\<close>
wneuper@59562
   124
                                                         
walther@59603
   125
ML \<open>Prog_Expr.eval_occurs_in\<close>
walther@59603
   126
ML \<open>@{thm last_thmI}\<close>
wneuper@59472
   127
ML \<open>@{thm Querkraft_Belastung}\<close>
neuper@41905
   128
wneuper@59472
   129
ML \<open>Celem.check_guhs_unique := false;\<close>
wneuper@59472
   130
ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
wneuper@59592
   131
ML \<open>@{theory "Isac_Knowledge"}\<close>
wneuper@59472
   132
ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
wneuper@59472
   133
  ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
neuper@42412
   134
wneuper@59472
   135
section \<open>State of approaching Isabelle by Isac\<close>
wneuper@59472
   136
text \<open>
neuper@55431
   137
  Mathias Lehnfeld gives the following list in his thesis in section 
neuper@55431
   138
  4.2.3 Relation to Ongoing Isabelle Development.
wneuper@59472
   139
\<close>
wneuper@59472
   140
subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
wneuper@59472
   141
text \<open>
neuper@55433
   142
  REPLACE BY KEStore... (has been overlooked)
neuper@55433
   143
    calcelems.sml:val rew_ord' = Unsynchronized.ref ...
neuper@55433
   144
  KEEP FOR TRACING
neuper@55433
   145
    calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
neuper@55433
   146
    calcelems.sml:val depth = Unsynchronized.ref 99999;
neuper@55433
   147
    calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
neuper@55433
   148
    calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
neuper@55433
   149
    Interpret/script.sml:val trace_script = Unsynchronized.ref false;
neuper@55433
   150
  KEEP FOR EASIER DEVELOPMENT
neuper@55433
   151
    calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
neuper@55433
   152
  KEEP FOR DEMOS
neuper@55433
   153
    Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
neuper@55433
   154
    Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
neuper@55433
   155
    Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
wneuper@59472
   156
\<close>
wneuper@59472
   157
subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
wneuper@59472
   158
subsection \<open>(2) Make Isac’s programming language usable\<close>
wneuper@59472
   159
subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
wneuper@59472
   160
text \<open>
neuper@55431
   161
  In 2002 isac already strived for floating point numbers. Since that time
neuper@55431
   162
  isac represents numerals as "Free", see below (*1*). These experiments are
neuper@55431
   163
  unsatisfactory with respect to logical soundness.
neuper@55431
   164
  Since Isabelle now has started to care about floating point numbers, it is high 
neuper@55431
   165
  time to adopt these together with the other numerals. Isabelle2012/13's numerals
neuper@55431
   166
  are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
neuper@55431
   167
  
neuper@55431
   168
  The transition from "Free" to standard numerals is a task to be scheduled for 
neuper@55431
   169
  several weeks. The urgency of this task follows from the experience, 
neuper@55431
   170
  that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
neuper@55431
   171
  some of the long identifiers of theorems which would tremendously simplify
neuper@55431
   172
  building a hierarchy of theorems according to (1.2), see (*2*) below.
wneuper@59472
   173
\<close>
wneuper@59472
   174
ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
neuper@55484
   175
wneuper@59472
   176
subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
wneuper@59472
   177
subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
neuper@55431
   178
neuper@37906
   179
end