src/Tools/isac/Build_Isac.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 13 Jan 2016 14:37:27 +0100
changeset 59206 ebf4a8a63371
parent 59191 4541844bab19
child 59259 d1c13ee4e1a2
permissions -rw-r--r--
cleanup theory imports, part 1

# Cleanup is triggered by failing session Protocol2015, which gives:
No such file: "/usr/local/isabisac/src/Tools/isac/Complex_Main.thy"
The error(s) above occurred for theory "Complex_Main"
(required by "Protocol" via "Build_Thydata" via "Isac" via "Frontend" via "xmlsrc" via "Interpret" via "ProgLang" via "Script" via "Tools" via "ListC" via "KEStore") (line 6 of "/usr/local/isabisac/src/Tools/isac/KEStore.thy")
No such file: "/usr/local/isabisac/src/Tools/isac/Knowledge/Real.thy"
:
:
# The newly introduced imports "~~/..." avoid the above kind of errors.
# However, these errors still remain:
Build started for Isabelle/Protocol2015 ...
Building Protocol2015 ...
Protocol2015: theory Classy
Protocol2015: theory Code_Generator
:
Protocol2015: theory Taylor
Protocol2015: theory Complex_Main
Protocol2015: theory KEStore
Protocol2015: theory ListC
Protocol2015 FAILED
*** Failed to load theory "RootRatEq" (unresolved "LinEq", "RatEq", "RootEq", "RootRat")
*** Failed to load theory "Partial_Fractions" (unresolved "RootRatEq")
:
:
wneuper@59178
     1
(*  Title:  build and test isac on Isabelle2014
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.
neuper@52062
     6
neuper@52062
     7
ATTENTION: no errors in this theory do not mean that there are no errors in Isac;
neuper@52062
     8
errors are rigorously detected when creating a heap.
neuper@37906
     9
*)
neuper@37906
    10
neuper@55276
    11
header {* Loading the isac mathengine;
neuper@55276
    12
  for debugging see text at begin (theory dependencies!) *}
neuper@37906
    13
neuper@48761
    14
theory Build_Isac 
wneuper@59206
    15
imports
wneuper@59182
    16
(* structure inherited from migration which began with Isabelle2009. improve?
wneuper@59182
    17
          theory KEStore
wneuper@59182
    18
            ML_file "~~/src/Tools/isac/library.sml"
wneuper@59182
    19
            ML_file "~~/src/Tools/isac/calcelems.sml"
wneuper@59182
    20
        theory ListC 
wneuper@59206
    21
          imports "~~/src/Tools/isac/KEStore"
wneuper@59182
    22
          ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
wneuper@59182
    23
          ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
wneuper@59182
    24
          ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
wneuper@59182
    25
      theory Tools imports ListC begin
wneuper@59182
    26
    theory Script imports Tools begin
wneuper@59182
    27
  theory ProgLang imports Script
wneuper@59182
    28
    ML_file "~~/src/Tools/isac/ProgLang/scrtools.sml"
neuper@52062
    29
*)        "ProgLang/ProgLang"
neuper@48761
    30
neuper@48763
    31
(* use    "Interpret/mstools.sml"
neuper@48763
    32
   use    "Interpret/ctree.sml"
neuper@48763
    33
   use    "Interpret/ptyps.sml"
wneuper@59106
    34
   use    "Interpret/generate.sml"                               
neuper@48763
    35
   use    "Interpret/calchead.sml"
neuper@48763
    36
   use    "Interpret/appl.sml"
neuper@48763
    37
   use    "Interpret/rewtools.sml"
neuper@48763
    38
   use    "Interpret/script.sml"
neuper@48763
    39
   use    "Interpret/solve.sml"
neuper@48763
    40
   use    "Interpret/inform.sml"
neuper@48763
    41
   use    "Interpret/mathengine.sml"
neuper@52062
    42
*)        "Interpret/Interpret"
neuper@48761
    43
neuper@52062
    44
(* use    "xmlsrc/mathml.sml"
neuper@48762
    45
   use    "xmlsrc/datatypes.sml"
neuper@41905
    46
   use    "xmlsrc/pbl-met-hierarchy.sml"
neuper@48895
    47
   use    "xmlsrc/thy-hierarchy.sml"                               
neuper@41905
    48
   use    "xmlsrc/interface-xml.sml"
neuper@52062
    49
*)        "xmlsrc/xmlsrc"
neuper@41905
    50
neuper@48763
    51
(* use     "Frontend/messages.sml"
neuper@41905
    52
   use     "Frontend/states.sml"
neuper@41905
    53
   use     "Frontend/interface.sml"
neuper@41905
    54
neuper@41905
    55
   use     "print_exn_G.sml"
neuper@52062
    56
*)         "Frontend/Frontend"
wneuper@59147
    57
wneuper@59147
    58
           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
wneuper@59147
    59
wneuper@59146
    60
           (* the Protocol for the connection isac-java -- Isabelle/Isac is built separately:
wneuper@59146
    61
              libisabelle$ /usr/local/isabisac/bin/isabelle build -D . -bv
wneuper@59148
    62
              here we partly check consistency libisabelle/Isac -- Isabelle/Isac: *)
wneuper@59191
    63
           (*"~~/libisabelle-protocol/isabelle-2015/Protocol"*)
neuper@48763
    64
begin
neuper@48760
    65
neuper@55276
    66
text {* 
neuper@55276
    67
  show theory dependencies using the graph browser, 
neuper@55276
    68
  open "browser_info/HOL/Isac/session.graph"
neuper@55276
    69
  and proceed from the ancestors towards the siblings.
neuper@55276
    70
*}
neuper@55276
    71
neuper@55431
    72
section {*check presence of definitions from directories*}
neuper@48763
    73
wneuper@59152
    74
ML {*
wneuper@59152
    75
*} ML {*
wneuper@59152
    76
*}
neuper@41905
    77
ML {* is_reall_dsc; (*from "ProgLang/scrtools.sml" *) *}
neuper@41905
    78
ML {* me; (*from "Interpret/mathengine.sml"*) *}
neuper@41905
    79
ML {* contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*) *}
neuper@41905
    80
ML {* print_exn_unit *}
neuper@48880
    81
ML {* list_rls (*from Atools.thy WN130615??? or ProgLang???*) *}
neuper@48763
    82
neuper@41905
    83
ML {* eval_occurs_in (*from Atools.thy*) *}
neuper@41905
    84
ML {* @{thm last_thmI} (*from Atools.thy*) *}
neuper@41931
    85
ML {*@{thm Querkraft_Belastung}*}
neuper@41905
    86
neuper@38009
    87
ML {* check_guhs_unique := false; *}
neuper@48880
    88
ML {* writeln "**** isac kernel = math-engine + Knowledge complete ******" *}
neuper@41943
    89
ML {* @{theory "Isac"} *}
neuper@55454
    90
ML {* (*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
neuper@55454
    91
  ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*) *}
neuper@42412
    92
neuper@55431
    93
section {* State of approaching Isabelle by Isac *}
neuper@55431
    94
text {* 
neuper@55431
    95
  Mathias Lehnfeld gives the following list in his thesis in section 
neuper@55431
    96
  4.2.3 Relation to Ongoing Isabelle Development.
neuper@55431
    97
*}
neuper@55431
    98
subsection {* (0) Survey on remaining Unsynchronized.ref *}
neuper@55433
    99
text {*
neuper@55433
   100
  REPLACE BY KEStore... (has been overlooked)
neuper@55433
   101
    calcelems.sml:val rew_ord' = Unsynchronized.ref ...
neuper@55433
   102
  KEEP FOR TRACING
neuper@55433
   103
    calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
neuper@55433
   104
    calcelems.sml:val depth = Unsynchronized.ref 99999;
neuper@55433
   105
    calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
neuper@55433
   106
    calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
neuper@55433
   107
    Interpret/script.sml:val trace_script = Unsynchronized.ref false;
neuper@55433
   108
  KEEP FOR EASIER DEVELOPMENT
neuper@55433
   109
    calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
neuper@55433
   110
  KEEP FOR DEMOS
neuper@55433
   111
    Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
neuper@55433
   112
    Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
neuper@55433
   113
    Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
neuper@55433
   114
*}
neuper@55431
   115
subsection {* (1) Exploit parallelism for concurrent session management *}
neuper@55431
   116
subsection {* (2) Make Isac’s programming language usable *}
neuper@55431
   117
subsection {* (3) Adopt Isabelle’s numeral computation for Isac *}
neuper@55431
   118
text {* 
neuper@55431
   119
  In 2002 isac already strived for floating point numbers. Since that time
neuper@55431
   120
  isac represents numerals as "Free", see below (*1*). These experiments are
neuper@55431
   121
  unsatisfactory with respect to logical soundness.
neuper@55431
   122
  Since Isabelle now has started to care about floating point numbers, it is high 
neuper@55431
   123
  time to adopt these together with the other numerals. Isabelle2012/13's numerals
neuper@55431
   124
  are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
neuper@55431
   125
  
neuper@55431
   126
  The transition from "Free" to standard numerals is a task to be scheduled for 
neuper@55431
   127
  several weeks. The urgency of this task follows from the experience, 
neuper@55431
   128
  that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
neuper@55431
   129
  some of the long identifiers of theorems which would tremendously simplify
neuper@55431
   130
  building a hierarchy of theorems according to (1.2), see (*2*) below.
neuper@55431
   131
*}
neuper@55431
   132
ML {*(*1*) Free ("123.456", HOLogic.realT) *}
neuper@55484
   133
neuper@55431
   134
subsection {* (4) Improve the efficiency of Isac’s rewrite-engine *}
neuper@55431
   135
subsection {* (5) Adopt Isabelle/jEdit for Isac *}
neuper@55431
   136
neuper@37906
   137
end