src/Tools/isac/Knowledge/Typefix.thy
author Walther Neuper <neuper@ist.tugraz.at>
Wed, 25 Aug 2010 16:20:07 +0200
branchisac-update-Isa09-2
changeset 37947 22235e4dbe5f
parent 37945 src/Tools/isac/IsacKnowledge/Typefix.thy@b63c897f4977
child 37948 ed85f172569c
permissions -rw-r--r--
renamed isac's directories and Build_Isac.thy

Scripts --> ProgLang
ME --> Interpret
IsacKnowledge --> Knowledge
neuper@37944
     1
(* Title:  fixed type for _RE_parsing of strings from frontend 
neuper@37944
     2
   Author: Walther Neuper
neuper@37944
     3
   9911xx 
neuper@37944
     4
   (c) due to copyright terms
neuper@37944
     5
   with hints from Markus Wenzel
neuper@37906
     6
 *)
neuper@37906
     7
neuper@37947
     8
theory Typefix imports "../ProgLang/Script" begin
neuper@37906
     9
neuper@37906
    10
syntax
neuper@37906
    11
       
neuper@37906
    12
  "_plus"  :: 'a
neuper@37906
    13
  "_minus" :: 'a
neuper@37906
    14
  "_umin"  :: 'a
neuper@37906
    15
  "_times" :: 'a
neuper@37906
    16
neuper@37906
    17
translations
neuper@37906
    18
neuper@37906
    19
  "op +"  => "_plus  :: [real, real]  => real"  (*infixl 65    *)
neuper@37906
    20
  "op -"  => "_minus :: [real, real] => real"   (*infixl 65    *)
neuper@37906
    21
  "uminus"=> "_umin  :: [real] => real"         (*"- _" [80] 80*)
neuper@37906
    22
  "op *"  => "_times :: [real, real] => real"   (*infixl 70    *)
neuper@37906
    23
neuper@37943
    24
ML {*
neuper@37906
    25
val parse_translation = 
neuper@37906
    26
    [("_plus", curry Term.list_comb (Syntax.const "op +")),  
neuper@37906
    27
     ("_minus", curry Term.list_comb (Syntax.const "op -")), 
neuper@37906
    28
     ("_umin", curry Term.list_comb (Syntax.const "uminus")),
neuper@37906
    29
     ("_times", curry Term.list_comb (Syntax.const "op *"))];
neuper@37943
    30
*}
neuper@37943
    31
neuper@37943
    32
end