blob: 997efd969fb389117507a0b3073005f202aa40fd (
plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
|
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2013 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Util
open Names
open Esubst
open Term
open Declarations
open Pre_env
open Nativevalues
(** This file defines the lambda code generation phase of the native compiler *)
type lambda =
| Lrel of name * int
| Lvar of identifier
| Lprod of lambda * lambda
| Llam of name array * lambda
| Llet of name * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of string * constant (* prefix, constant name *)
| Lcase of annot_sw * lambda * lambda * lam_branches
(* annotations, term being matched, accu, branches *)
| Lif of lambda * lambda * lambda
| Lfix of (int array * int) * fix_decl
| Lcofix of int * fix_decl
| Lmakeblock of string * constructor * int * lambda array
(* prefix, constructor name, constructor tag, arguments *)
(* A fully applied constructor *)
| Lconstruct of string * constructor (* prefix, constructor name *)
(* A partially applied constructor *)
| Lval of Nativevalues.t
| Lsort of sorts
| Lind of string * inductive (* prefix, inductive name *)
| Llazy
| Lforce
and lam_branches = (constructor * name array * lambda) array
and fix_decl = name array * lambda array * lambda array
val decompose_Llam : lambda -> Names.name array * lambda
val decompose_Llam_Llet : lambda -> (Names.name * lambda option) array * lambda
val is_lazy : constr -> bool
val mk_lazy : lambda -> lambda
val get_allias : env -> constant -> constant
val lambda_of_constr : env -> types -> lambda
|