src/Tools/isac/Build_Isac.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 21 Nov 2018 12:32:54 +0100
changeset 59472 3e904f8ec16c
parent 59470 e11233d9b98e
child 59486 141c89a20197
permissions -rw-r--r--
update to new Isabelle conventions: {*...*} to \<open>...\<close>
wneuper@59459
     1
(*  Title:  build the Isac Mathengine
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.
wneuper@59344
     6
For debugging see text at begin (theory dependencies!)
neuper@52062
     7
neuper@52062
     8
ATTENTION: no errors in this theory do not mean that there are no errors in Isac;
neuper@52062
     9
errors are rigorously detected when creating a heap.
neuper@37906
    10
*)
neuper@37906
    11
wneuper@59347
    12
theory Build_Isac 
wneuper@59206
    13
imports
wneuper@59182
    14
(* structure inherited from migration which began with Isabelle2009. improve?
wneuper@59182
    15
          theory KEStore
wneuper@59182
    16
            ML_file "~~/src/Tools/isac/library.sml"
wneuper@59182
    17
            ML_file "~~/src/Tools/isac/calcelems.sml"
wneuper@59182
    18
        theory ListC 
wneuper@59206
    19
          imports "~~/src/Tools/isac/KEStore"
wneuper@59182
    20
          ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
wneuper@59182
    21
          ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
wneuper@59182
    22
          ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
wneuper@59182
    23
      theory Tools imports ListC begin
wneuper@59182
    24
    theory Script imports Tools begin
wneuper@59182
    25
  theory ProgLang imports Script
wneuper@59182
    26
    ML_file "~~/src/Tools/isac/ProgLang/scrtools.sml"
neuper@52062
    27
*)        "ProgLang/ProgLang"
neuper@48761
    28
neuper@48763
    29
(* use    "Interpret/mstools.sml"
neuper@48763
    30
   use    "Interpret/ctree.sml"
neuper@48763
    31
   use    "Interpret/ptyps.sml"
wneuper@59106
    32
   use    "Interpret/generate.sml"                               
neuper@48763
    33
   use    "Interpret/calchead.sml"
neuper@48763
    34
   use    "Interpret/appl.sml"
neuper@48763
    35
   use    "Interpret/rewtools.sml"
neuper@48763
    36
   use    "Interpret/script.sml"
neuper@48763
    37
   use    "Interpret/solve.sml"
neuper@48763
    38
   use    "Interpret/inform.sml"
neuper@48763
    39
   use    "Interpret/mathengine.sml"
neuper@52062
    40
*)        "Interpret/Interpret"
neuper@48761
    41
neuper@52062
    42
(* use    "xmlsrc/mathml.sml"
neuper@48762
    43
   use    "xmlsrc/datatypes.sml"
neuper@41905
    44
   use    "xmlsrc/pbl-met-hierarchy.sml"
neuper@48895
    45
   use    "xmlsrc/thy-hierarchy.sml"                               
neuper@41905
    46
   use    "xmlsrc/interface-xml.sml"
neuper@52062
    47
*)        "xmlsrc/xmlsrc"
neuper@41905
    48
neuper@48763
    49
(* use     "Frontend/messages.sml"
neuper@41905
    50
   use     "Frontend/states.sml"
neuper@41905
    51
   use     "Frontend/interface.sml"
neuper@41905
    52
neuper@41905
    53
   use     "print_exn_G.sml"
neuper@52062
    54
*)         "Frontend/Frontend"
wneuper@59147
    55
wneuper@59147
    56
           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
wneuper@59147
    57
wneuper@59146
    58
           (* the Protocol for the connection isac-java -- Isabelle/Isac is built separately:
wneuper@59148
    59
              here we partly check consistency libisabelle/Isac -- Isabelle/Isac: *)
wneuper@59470
    60
           Protocol.Protocol
wneuper@59469
    61
neuper@48763
    62
begin
neuper@48760
    63
wneuper@59472
    64
text \<open>
neuper@55276
    65
  show theory dependencies using the graph browser, 
neuper@55276
    66
  open "browser_info/HOL/Isac/session.graph"
neuper@55276
    67
  and proceed from the ancestors towards the siblings.
wneuper@59472
    68
\<close>
neuper@55276
    69
wneuper@59472
    70
section \<open>check presence of definitions from directories\<close>
neuper@48763
    71
wneuper@59263
    72
(*declare [[ML_print_depth = 999]]*)
wneuper@59472
    73
ML \<open>
wneuper@59472
    74
\<close> ML \<open>
wneuper@59472
    75
\<close>
wneuper@59382
    76
wneuper@59382
    77
(*==============================================================================================
wneuper@59382
    78
  The test below from calculate.sml raises an exception with the cleaned Rewrite.
wneuper@59382
    79
  The differences to the working 'fun rewrite' are that significant,
wneuper@59382
    80
  that we want to rely on 'hg revert'.
wneuper@59382
    81
  For that purpose we save this changeset with a broken 'fun rewrite' and tests not running.
wneuper@59382
    82
*)
wneuper@59472
    83
ML \<open>
wneuper@59472
    84
\<close> ML \<open>
wneuper@59382
    85
(*--------------(2): does divide work in Test_simplify ?: ------*)
wneuper@59382
    86
 val thy = @{theory Test};
wneuper@59389
    87
 val t = (Thm.term_of o the o (TermC.parse thy)) "6 / 2";
wneuper@59382
    88
 val rls = Test_simplify;
wneuper@59472
    89
\<close> ML \<open>
wneuper@59386
    90
 val (t,_) = the (Rewrite.rewrite_set_ thy false rls t);
wneuper@59382
    91
(*val t = Free ("3","Real.real") : term*)
wneuper@59472
    92
\<close> ML \<open>
wneuper@59382
    93
wneuper@59382
    94
(*--------------(3): is_const works ?: -------------------------------------*)
wneuper@59389
    95
 val t = (Thm.term_of o the o (TermC.parse @{theory})) "2 is_const";
wneuper@59472
    96
\<close> ML \<open>
wneuper@59386
    97
 Rewrite.rewrite_set_ @{theory Test} false tval_rls t;
wneuper@59382
    98
wneuper@59383
    99
(* exception TERM raised (line 270 of "~~/src/HOL/Tools/hologic.ML"):
wneuper@59382
   100
  dest_eq
wneuper@59382
   101
  0 \<le> ?a \<Longrightarrow> (0 \<le> sqrt ?a) = True*)
wneuper@59472
   102
\<close> ML \<open>
wneuper@59472
   103
\<close>
wneuper@59382
   104
(*==============================================================================================*)
wneuper@59382
   105
wneuper@59472
   106
ML \<open>Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
wneuper@59472
   107
ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
wneuper@59472
   108
ML \<open>LTool.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
wneuper@59472
   109
ML \<open>Math_Engine.me; (*from "Interpret/mathengine.sml"*)\<close>
wneuper@59472
   110
ML \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
wneuper@59472
   111
ML \<open>print_exn_unit\<close>
wneuper@59472
   112
ML \<open>list_rls (*from Atools.thy WN130615??? or ProgLang???*)\<close>
neuper@48763
   113
wneuper@59472
   114
ML \<open>eval_occurs_in (*from Atools.thy*)\<close>
wneuper@59472
   115
ML \<open>@{thm last_thmI} (*from Atools.thy*)\<close>
wneuper@59472
   116
ML \<open>@{thm Querkraft_Belastung}\<close>
neuper@41905
   117
wneuper@59472
   118
ML \<open>Celem.check_guhs_unique := false;\<close>
wneuper@59472
   119
ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
wneuper@59472
   120
ML \<open>@{theory "Isac"}\<close>
wneuper@59472
   121
ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
wneuper@59472
   122
  ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
neuper@42412
   123
wneuper@59472
   124
section \<open>State of approaching Isabelle by Isac\<close>
wneuper@59472
   125
text \<open>
neuper@55431
   126
  Mathias Lehnfeld gives the following list in his thesis in section 
neuper@55431
   127
  4.2.3 Relation to Ongoing Isabelle Development.
wneuper@59472
   128
\<close>
wneuper@59472
   129
subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
wneuper@59472
   130
text \<open>
neuper@55433
   131
  REPLACE BY KEStore... (has been overlooked)
neuper@55433
   132
    calcelems.sml:val rew_ord' = Unsynchronized.ref ...
neuper@55433
   133
  KEEP FOR TRACING
neuper@55433
   134
    calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
neuper@55433
   135
    calcelems.sml:val depth = Unsynchronized.ref 99999;
neuper@55433
   136
    calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
neuper@55433
   137
    calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
neuper@55433
   138
    Interpret/script.sml:val trace_script = Unsynchronized.ref false;
neuper@55433
   139
  KEEP FOR EASIER DEVELOPMENT
neuper@55433
   140
    calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
neuper@55433
   141
  KEEP FOR DEMOS
neuper@55433
   142
    Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
neuper@55433
   143
    Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
neuper@55433
   144
    Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
wneuper@59472
   145
\<close>
wneuper@59472
   146
subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
wneuper@59472
   147
subsection \<open>(2) Make Isac’s programming language usable\<close>
wneuper@59472
   148
subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
wneuper@59472
   149
text \<open>
neuper@55431
   150
  In 2002 isac already strived for floating point numbers. Since that time
neuper@55431
   151
  isac represents numerals as "Free", see below (*1*). These experiments are
neuper@55431
   152
  unsatisfactory with respect to logical soundness.
neuper@55431
   153
  Since Isabelle now has started to care about floating point numbers, it is high 
neuper@55431
   154
  time to adopt these together with the other numerals. Isabelle2012/13's numerals
neuper@55431
   155
  are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
neuper@55431
   156
  
neuper@55431
   157
  The transition from "Free" to standard numerals is a task to be scheduled for 
neuper@55431
   158
  several weeks. The urgency of this task follows from the experience, 
neuper@55431
   159
  that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
neuper@55431
   160
  some of the long identifiers of theorems which would tremendously simplify
neuper@55431
   161
  building a hierarchy of theorems according to (1.2), see (*2*) below.
wneuper@59472
   162
\<close>
wneuper@59472
   163
ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
neuper@55484
   164
wneuper@59472
   165
subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
wneuper@59472
   166
subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
neuper@55431
   167
neuper@37906
   168
end