| 13208 |      1 | (*  Title:    HOL/Prolog/Type.thy
 | 
|  |      2 |     ID:       $Id$
 | 
|  |      3 |     Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
 | 
|  |      4 | *)
 | 
| 9015 |      5 | 
 | 
| 17311 |      6 | header {* Type inference *}
 | 
| 9015 |      7 | 
 | 
| 17311 |      8 | theory Type
 | 
|  |      9 | imports Func
 | 
|  |     10 | begin
 | 
| 9015 |     11 | 
 | 
| 17311 |     12 | typedecl ty
 | 
| 9015 |     13 | 
 | 
| 17311 |     14 | consts
 | 
|  |     15 |   bool    :: ty
 | 
|  |     16 |   nat     :: ty
 | 
|  |     17 |   "->"    :: "ty => ty => ty"       (infixr 20)
 | 
|  |     18 |   typeof  :: "[tm, ty] => bool"
 | 
|  |     19 |   anyterm :: tm
 | 
| 9015 |     20 | 
 | 
| 17311 |     21 | axioms  common_typeof:   "
 | 
| 9015 |     22 | typeof (app M N) B       :- typeof M (A -> B) & typeof N A..
 | 
|  |     23 | 
 | 
|  |     24 | typeof (cond C L R) A :- typeof C bool & typeof L A & typeof R A..
 | 
|  |     25 | typeof (fix F)   A       :- (!x. typeof x A => typeof (F  x) A)..
 | 
|  |     26 | 
 | 
|  |     27 | typeof true  bool..
 | 
|  |     28 | typeof false bool..
 | 
|  |     29 | typeof (M and N) bool :- typeof M bool & typeof N bool..
 | 
|  |     30 | 
 | 
|  |     31 | typeof (M eq  N) bool :- typeof M T    & typeof N T   ..
 | 
|  |     32 | 
 | 
|  |     33 | typeof  Z    nat..
 | 
|  |     34 | typeof (S N) nat :- typeof N nat..
 | 
|  |     35 | typeof (M + N) nat :- typeof M nat & typeof N nat..
 | 
|  |     36 | typeof (M - N) nat :- typeof M nat & typeof N nat..
 | 
|  |     37 | typeof (M * N) nat :- typeof M nat & typeof N nat"
 | 
|  |     38 | 
 | 
| 17311 |     39 | axioms good_typeof:     "
 | 
| 9015 |     40 | typeof (abs Bo) (A -> B) :- (!x. typeof x A => typeof (Bo x) B)"
 | 
|  |     41 | 
 | 
| 17311 |     42 | axioms bad1_typeof:     "
 | 
| 9015 |     43 | typeof (abs Bo) (A -> B) :- (typeof varterm A => typeof (Bo varterm) B)"
 | 
|  |     44 | 
 | 
| 17311 |     45 | axioms bad2_typeof:     "
 | 
| 9015 |     46 | typeof (abs Bo) (A -> B) :- (typeof anyterm A => typeof (Bo anyterm) B)"
 | 
|  |     47 | 
 | 
| 17311 |     48 | ML {* use_legacy_bindings (the_context ()) *}
 | 
|  |     49 | 
 | 
| 9015 |     50 | end
 |