| 1478 |      1 | (*  Title:      ZF/Coind/Language.thy
 | 
|  |      2 |     Author:     Jacob Frost, Cambridge University Computer Laboratory
 | 
| 915 |      3 |     Copyright   1995  University of Cambridge
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
| 16417 |      6 | theory Language imports Main begin
 | 
| 915 |      7 | 
 | 
| 932 |      8 | consts
 | 
| 1478 |      9 |   Const :: i                    (* Abstract type of constants *)
 | 
| 12595 |     10 |   c_app :: "[i,i] => i"         (* Abstract constructor for fun application*)
 | 
|  |     11 | 
 | 
| 915 |     12 | 
 | 
| 12595 |     13 | text{*these really can't be definitions without losing the abstraction*}
 | 
|  |     14 | axioms
 | 
|  |     15 |   constNEE:  "c \<in> Const ==> c \<noteq> 0"
 | 
|  |     16 |   c_appI:    "[| c1 \<in> Const; c2 \<in> Const |] ==> c_app(c1,c2) \<in> Const"
 | 
| 915 |     17 | 
 | 
|  |     18 | 
 | 
|  |     19 | consts
 | 
| 1478 |     20 |   Exp   :: i                    (* Datatype of expressions *)
 | 
|  |     21 |   ExVar :: i                    (* Abstract type of variables *)
 | 
| 6112 |     22 | 
 | 
|  |     23 | datatype
 | 
| 12595 |     24 |   "Exp" = e_const ("c \<in> Const")
 | 
|  |     25 |         | e_var ("x \<in> ExVar")
 | 
|  |     26 |         | e_fn ("x \<in> ExVar","e \<in> Exp")
 | 
|  |     27 |         | e_fix ("x1 \<in> ExVar","x2 \<in> ExVar","e \<in> Exp")
 | 
|  |     28 |         | e_app ("e1 \<in> Exp","e2 \<in> Exp")
 | 
| 915 |     29 | 
 | 
|  |     30 | end
 |