src/HOL/MiniML/MiniML.thy
author berghofe
Fri, 24 Jul 1998 13:19:38 +0200
changeset 5184 9b8547a9496a
parent 4502 337c073de95e
child 14422 b8da5f258b04
permissions -rw-r--r--
Adapted to new datatype package.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     1
(* Title:     HOL/MiniML/MiniML.thy
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     2
   ID:        $Id$
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
     3
   Author:    Wolfgang Naraschewski and Tobias Nipkow
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
     4
   Copyright  1996 TU Muenchen
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     5
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     6
Mini_ML with type inference rules.
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     7
*)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     8
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
     9
MiniML = Generalize + 
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    10
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    11
(* expressions *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    12
datatype
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    13
        expr = Var nat | Abs expr | App expr expr | LET expr expr
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    14
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    15
(* type inference rules *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    16
consts
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    17
        has_type :: "(ctxt * expr * typ)set"
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    18
syntax
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    19
        "@has_type" :: [ctxt, expr, typ] => bool
1525
d127436567d0 modified priorities in syntax
nipkow
parents: 1476
diff changeset
    20
                       ("((_) |-/ (_) :: (_))" [60,0,60] 60)
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    21
translations 
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    22
        "A |- e :: t" == "(A,e,t):has_type"
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    23
1790
2f3694c50101 Quotes now optional around inductive set
paulson
parents: 1525
diff changeset
    24
inductive has_type
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    25
intrs
4502
337c073de95e nth -> !
nipkow
parents: 2525
diff changeset
    26
        VarI "[| n < length A; t <| A!n |] ==> A |- Var n :: t"
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    27
        AbsI "[| (mk_scheme t1)#A |- e :: t2 |] ==> A |- Abs e :: t1 -> t2"
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    28
        AppI "[| A |- e1 :: t2 -> t1; A |- e2 :: t2 |] 
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    29
              ==> A |- App e1 e2 :: t1"
477c05586286 The new version of MiniML including "let".
nipkow
parents: 1790
diff changeset
    30
        LETI "[| A |- e1 :: t1; (gen A t1)#A |- e2 :: t |] ==> A |- LET e1 e2 :: t"
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    31
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    32
end