915
|
1 |
(* Title: ZF/Coind/Language.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Jacob Frost, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1995 University of Cambridge
|
|
5 |
*)
|
|
6 |
|
|
7 |
Language ="Datatype" + QUniv +
|
|
8 |
|
|
9 |
(* Abstract type of constants *)
|
|
10 |
|
|
11 |
consts
|
|
12 |
Const :: "i"
|
|
13 |
rules
|
|
14 |
constU "c:Const ==> c:univ(A)"
|
|
15 |
constNEE "c:Const ==> c ~= 0"
|
|
16 |
|
|
17 |
(* A function that captures how constant functions are applied to
|
|
18 |
constants *)
|
|
19 |
|
|
20 |
consts
|
|
21 |
c_app :: "[i,i] => i"
|
|
22 |
rules
|
|
23 |
c_appI "[| c1:Const; c2:Const |] ==> c_app(c1,c2):Const"
|
|
24 |
|
|
25 |
|
|
26 |
(* Abstract type of variables *)
|
|
27 |
|
|
28 |
consts
|
|
29 |
ExVar :: "i"
|
|
30 |
rules
|
|
31 |
exvarU "x:ExVar ==> x:univ(A)"
|
|
32 |
|
|
33 |
|
|
34 |
(* Datatype of expressions *)
|
|
35 |
|
|
36 |
consts
|
|
37 |
Exp :: "i"
|
|
38 |
datatype
|
|
39 |
"Exp" =
|
|
40 |
e_const("c:Const") |
|
|
41 |
e_var("x:ExVar") |
|
|
42 |
e_fn("x:ExVar","e:Exp") |
|
|
43 |
e_fix("x1:ExVar","x2:ExVar","e:Exp") |
|
|
44 |
e_app("e1:Exp","e2:Exp")
|
|
45 |
type_intrs "[constU,exvarU]"
|
|
46 |
|
|
47 |
end
|
|
48 |
|
|
49 |
|
|
50 |
|
|
51 |
|