aboutsummaryrefslogtreecommitdiff
path: root/kernel/cPrimitives.ml
diff options
context:
space:
mode:
authorPierre Roux2019-04-02 22:39:32 +0200
committerPierre Roux2019-11-01 10:19:59 +0100
commitf93684a412f067622a5026c406bc76032c30b6e9 (patch)
tree94965ae5e5d454b0ebb0d4266dd8a27f5487ddf3 /kernel/cPrimitives.ml
parent6694a1811dc4e961a81fb4464cf5aaf05f1b5752 (diff)
Declare type of primitives in CPrimitives
Rather than in typeops
Diffstat (limited to 'kernel/cPrimitives.ml')
-rw-r--r--kernel/cPrimitives.ml72
1 files changed, 38 insertions, 34 deletions
diff --git a/kernel/cPrimitives.ml b/kernel/cPrimitives.ml
index d854cadd15..d433cdc1ba 100644
--- a/kernel/cPrimitives.ml
+++ b/kernel/cPrimitives.ml
@@ -90,58 +90,62 @@ let to_string = function
| Int63le -> "le"
| Int63compare -> "compare"
-type arg_kind =
- | Kparam (* not needed for the evaluation of the primitive when it reduces *)
- | Kwhnf (* need to be reduced in whnf before reducing the primitive *)
- | Karg (* no need to be reduced in whnf. example: [v] in [Array.set t i v] *)
+type prim_type =
+ | PT_int63
-type args_red = arg_kind list
+type 'a prim_ind =
+ | PIT_bool : unit prim_ind
+ | PIT_carry : prim_type prim_ind
+ | PIT_pair : (prim_type * prim_type) prim_ind
+ | PIT_cmp : unit prim_ind
-(* Invariant only argument of type int63 or an inductive can
- have kind Kwhnf *)
+type prim_ind_ex = PIE : 'a prim_ind -> prim_ind_ex
-let kind = function
- | Int63head0 | Int63tail0 -> [Kwhnf]
+type ind_or_type =
+ | PITT_ind : 'a prim_ind * 'a -> ind_or_type
+ | PITT_type : prim_type -> ind_or_type
+let types =
+ let int_ty = PITT_type PT_int63 in
+ function
+ | Int63head0 | Int63tail0 -> [int_ty; int_ty]
| Int63add | Int63sub | Int63mul
| Int63div | Int63mod
| Int63lsr | Int63lsl
- | Int63land | Int63lor | Int63lxor
- | Int63addc | Int63subc
- | Int63addCarryC | Int63subCarryC | Int63mulc | Int63diveucl
- | Int63eq | Int63lt | Int63le | Int63compare -> [Kwhnf; Kwhnf]
+ | Int63land | Int63lor | Int63lxor -> [int_ty; int_ty; int_ty]
+ | Int63addc | Int63subc | Int63addCarryC | Int63subCarryC ->
+ [int_ty; int_ty; PITT_ind (PIT_carry, PT_int63)]
+ | Int63mulc | Int63diveucl ->
+ [int_ty; int_ty; PITT_ind (PIT_pair, (PT_int63, PT_int63))]
+ | Int63eq | Int63lt | Int63le -> [int_ty; int_ty; PITT_ind (PIT_bool, ())]
+ | Int63compare -> [int_ty; int_ty; PITT_ind (PIT_cmp, ())]
+ | Int63div21 ->
+ [int_ty; int_ty; int_ty; PITT_ind (PIT_pair, (PT_int63, PT_int63))]
+ | Int63addMulDiv -> [int_ty; int_ty; int_ty; int_ty]
- | Int63div21 | Int63addMulDiv -> [Kwhnf; Kwhnf; Kwhnf]
+type arg_kind =
+ | Kparam (* not needed for the evaluation of the primitive when it reduces *)
+ | Kwhnf (* need to be reduced in whnf before reducing the primitive *)
+ | Karg (* no need to be reduced in whnf. example: [v] in [Array.set t i v] *)
-let arity = function
- | Int63head0 | Int63tail0 -> 1
- | Int63add | Int63sub | Int63mul
- | Int63div | Int63mod
- | Int63lsr | Int63lsl
- | Int63land | Int63lor | Int63lxor
- | Int63addc | Int63subc
- | Int63addCarryC | Int63subCarryC | Int63mulc | Int63diveucl
- | Int63eq | Int63lt | Int63le
- | Int63compare -> 2
+type args_red = arg_kind list
- | Int63div21 | Int63addMulDiv -> 3
+(* Invariant only argument of type int63 or an inductive can
+ have kind Kwhnf *)
-(** Special Entries for Register **)
+let arity t = List.length (types t) - 1
-type prim_ind =
- | PIT_bool
- | PIT_carry
- | PIT_pair
- | PIT_cmp
+let kind t =
+ let rec aux n = if n <= 0 then [] else Kwhnf :: aux (n - 1) in
+ aux (arity t)
-type prim_type =
- | PT_int63
+(** Special Entries for Register **)
type op_or_type =
| OT_op of t
| OT_type of prim_type
-let prim_ind_to_string = function
+let prim_ind_to_string (type a) (p : a prim_ind) = match p with
| PIT_bool -> "bool"
| PIT_carry -> "carry"
| PIT_pair -> "pair"