a5671964fc86f5486b5600682c6af0acb962da19
[coq-hetmet.git] / src / HaskCoreTypes.v
1 (*********************************************************************************************************************************)
2 (* HaskCoreTypes: basically GHC's TypeRep.Type imported into Coqland                                                             *)
3 (*********************************************************************************************************************************)
4
5 Generalizable All Variables.
6 Require Import Preamble.
7 Require Import General.
8 Require Import Coq.Strings.String.
9 Require Import HaskGeneral.
10 Require Import HaskLiterals.
11 Require Import HaskCoreVars.
12
13 Variable CoreName           : Type.                                               Extract Inlined Constant CoreName => "Name.Name".
14 Variable coreName_eq        : forall (a b:CoreName),   sumbool (a=b) (not (a=b)). Extract Inlined Constant coreName_eq => "(==)".
15 Axiom    coreName_eq_refl   : ∀ v, (coreName_eq v v)=(left _ (refl_equal v)).
16 Instance CoreNameEqDecidable : EqDecidable CoreName :=
17 { eqd_dec            := coreName_eq
18 }.
19
20 Inductive CoreIPName        : Type -> Type := .                               Extract Inductive CoreIPName => "CoreSyn.IPName" [ ].
21
22 (* this exracts onto TypeRep.Type, on the nose *)
23 Inductive CoreType :=
24 | TyVarTy  :             CoreVar                   -> CoreType
25 | AppTy    :             CoreType ->      CoreType -> CoreType   (* first arg must be AppTy or TyVarTy*)
26 | TyConApp : forall {n}, TyCon n  -> list CoreType -> CoreType
27 | FunTy    :             CoreType ->      CoreType -> CoreType   (* technically redundant since we have FunTyCon *)
28 | ForAllTy :             CoreVar  ->      CoreType -> CoreType
29 | PredTy   :             PredType                  -> CoreType
30 with PredType :=
31 | ClassP   : forall {n}, Class_ n            -> list CoreType -> PredType
32 | IParam   :             CoreIPName CoreName -> CoreType      -> PredType
33 | EqPred   :             CoreType            -> CoreType      -> PredType.
34 Extract Inductive CoreType =>
35    "TypeRep.Type" [ "TypeRep.TyVarTy" "TypeRep.AppTy" "TypeRep.TyConApp" "TypeRep.FunTy" "TypeRep.ForAllTy" "TypeRep.PredTy" ].
36 Extract Inductive PredType =>
37    "TypeRep.PredType" [ "TypeRep.ClassP" "TypeRep.IParam" "TypeRep.EqPred" ].
38
39 Variable coreTypeToString      : CoreType     -> string.    Extract Inlined Constant coreTypeToString       => "outputableToString".
40 Variable coreNameToString      : CoreName     -> string.    Extract Inlined Constant coreNameToString       => "outputableToString".
41
42 Variable CoreCoercion          : Type.                      Extract Inlined Constant CoreCoercion           => "Coercion.Coercion".
43 Variable coreCoercionToString  : CoreCoercion -> string.    Extract Inlined Constant coreCoercionToString   => "outputableToString".
44 Variable coreCoercionKind      : CoreCoercion -> CoreType*CoreType.
45  Extract Inlined Constant coreCoercionKind => "Coercion.coercionKind".
46
47 Variable coretype_eq_dec : forall (c1 c2:CoreType), sumbool (eq c1 c2) (not (eq c1 c2)).
48   Extract Inlined Constant coretype_eq_dec => "Type.coreEqType".
49   Instance CoreTypeEqDecidable : EqDecidable CoreType.
50     apply Build_EqDecidable.
51     apply coretype_eq_dec.
52     Defined.
53
54 Variable kindOfCoreType : CoreType -> Kind.      Extract Inlined Constant kindOfCoreType => "(coreKindToKind . Coercion.typeKind)".