| author | wenzelm | 
| Fri, 15 Feb 2002 20:41:19 +0100 | |
| changeset 12892 | a0b2acb7d6fa | 
| parent 12338 | de0f4a63baa5 | 
| child 13208 | 965f95a3abd9 | 
| permissions | -rw-r--r-- | 
| 9015 | 1 | (* type inference *) | 
| 2 | ||
| 3 | Type = Func + | |
| 4 | ||
| 5 | types ty | |
| 6 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
9015diff
changeset | 7 | arities ty :: type | 
| 9015 | 8 | |
| 9 | consts bool :: ty | |
| 10 | nat :: ty | |
| 11 | "->" :: ty => ty => ty (infixr 20) | |
| 12 | typeof :: [tm, ty] => bool | |
| 13 | anyterm :: tm | |
| 14 | ||
| 15 | rules common_typeof " | |
| 16 | typeof (app M N) B :- typeof M (A -> B) & typeof N A.. | |
| 17 | ||
| 18 | typeof (cond C L R) A :- typeof C bool & typeof L A & typeof R A.. | |
| 19 | typeof (fix F) A :- (!x. typeof x A => typeof (F x) A).. | |
| 20 | ||
| 21 | typeof true bool.. | |
| 22 | typeof false bool.. | |
| 23 | typeof (M and N) bool :- typeof M bool & typeof N bool.. | |
| 24 | ||
| 25 | typeof (M eq N) bool :- typeof M T & typeof N T .. | |
| 26 | ||
| 27 | typeof Z nat.. | |
| 28 | typeof (S N) nat :- typeof N nat.. | |
| 29 | typeof (M + N) nat :- typeof M nat & typeof N nat.. | |
| 30 | typeof (M - N) nat :- typeof M nat & typeof N nat.. | |
| 31 | typeof (M * N) nat :- typeof M nat & typeof N nat" | |
| 32 | ||
| 33 | rules good_typeof " | |
| 34 | typeof (abs Bo) (A -> B) :- (!x. typeof x A => typeof (Bo x) B)" | |
| 35 | ||
| 36 | rules bad1_typeof " | |
| 37 | typeof (abs Bo) (A -> B) :- (typeof varterm A => typeof (Bo varterm) B)" | |
| 38 | ||
| 39 | rules bad2_typeof " | |
| 40 | typeof (abs Bo) (A -> B) :- (typeof anyterm A => typeof (Bo anyterm) B)" | |
| 41 | ||
| 42 | end |