src/HOL/MiniML/MiniML.ML
author nipkow
Wed, 25 Oct 1995 09:46:46 +0100
changeset 1300 c7a8f374339b
child 1521 4ed3004ff75e
permissions -rw-r--r--
New theory: type inference for let-free MiniML
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.ML
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     2
   ID:        $Id$
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     3
   Author:    Dieter Nazareth and Tobias Nipkow
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     4
   Copyright  1995 TU Muenchen
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
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     7
open MiniML;
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     8
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     9
Addsimps has_type.intrs;
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    10
Addsimps [Un_upper1,Un_upper2];
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    11
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    12
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    13
(* has_type is closed w.r.t. substitution *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    14
goal MiniML.thy
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    15
     "!a e t. (a |- e :: t) --> ($ s a |- e :: $ s t)";
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    16
by (rtac has_type.mutual_induct 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    17
(* case VarI *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    18
by (asm_full_simp_tac (!simpset addsimps [app_subst_list]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    19
by (forw_inst_tac [("f1","$ s")] (nth_map RS sym) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    20
by( fast_tac (HOL_cs addIs [has_type.VarI] addss (!simpset delsimps [nth_map])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    21
(* case AbsI *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    22
by (asm_full_simp_tac (!simpset addsimps [app_subst_list]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    23
(* case AppI *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    24
by (Asm_full_simp_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    25
bind_thm ("has_type_cl_sub", result() RS spec RS spec RS spec RS mp);