src/Tools/isac/Knowledge/Typefix.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 26 Aug 2010 10:03:53 +0200
branchisac-update-Isa09-2
changeset 37949 aaf528d3ebd5
parent 37948 ed85f172569c
child 37959 cc24d0f70544
permissions -rw-r--r--
messed file dependencies.

The mess becomes apparent when unduly simplifying to
'theory Atools imports Complex_Main'
(*theory Atools imports Descript Typefix begin ...WOULD BE REQUIRED,
but hangs -- this needs to be clarified as well*):
The simplified header + use/s "../ProgLang/term.sml" removed:
*** Error (line 127 of "/usr/local/isabisac/src/Tools/isac/Knowledge/Atools.thy"):
*** Value or constructor (vars) has not been declared
but now arrive at:
*** Error (line 26 of "/usr/local/isabisac/src/Tools/isac/ProgLang/term.sml"):
*** Value or constructor (indent) has not been declared
which is definitely a mess, because ProgLang/term.sml appeared
already earlier in Build_Isac.thy
neuper@37944
     1
(* Title:  fixed type for _RE_parsing of strings from frontend 
neuper@37948
     2
   Author: Walther Neuper, hints from Makarius Wenzel
neuper@37944
     3
   9911xx 
neuper@37944
     4
   (c) due to copyright terms
neuper@37906
     5
 *)
neuper@37906
     6
neuper@37949
     7
theory Typefix imports "../ProgLang/Script" begin
neuper@37906
     8
neuper@37906
     9
syntax
neuper@37906
    10
       
neuper@37906
    11
  "_plus"  :: 'a
neuper@37906
    12
  "_minus" :: 'a
neuper@37906
    13
  "_umin"  :: 'a
neuper@37906
    14
  "_times" :: 'a
neuper@37906
    15
neuper@37906
    16
translations
neuper@37906
    17
neuper@37906
    18
  "op +"  => "_plus  :: [real, real]  => real"  (*infixl 65    *)
neuper@37906
    19
  "op -"  => "_minus :: [real, real] => real"   (*infixl 65    *)
neuper@37906
    20
  "uminus"=> "_umin  :: [real] => real"         (*"- _" [80] 80*)
neuper@37906
    21
  "op *"  => "_times :: [real, real] => real"   (*infixl 70    *)
neuper@37906
    22
neuper@37943
    23
ML {*
neuper@37906
    24
val parse_translation = 
neuper@37906
    25
    [("_plus", curry Term.list_comb (Syntax.const "op +")),  
neuper@37906
    26
     ("_minus", curry Term.list_comb (Syntax.const "op -")), 
neuper@37906
    27
     ("_umin", curry Term.list_comb (Syntax.const "uminus")),
neuper@37906
    28
     ("_times", curry Term.list_comb (Syntax.const "op *"))];
neuper@37943
    29
*}
neuper@37943
    30
neuper@37943
    31
end