X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskCoreVars.v;h=d158f050e07a299215f47a1f01ed664bf5f0abf5;hp=562b478ff99ae0cb2da1e0777e97851ac87ba1ab;hb=2ec43bc871b579bac89707988c4855ee1d6c8eda;hpb=24445b56cb514694c603c342d77cbc8329a4b0aa diff --git a/src/HaskCoreVars.v b/src/HaskCoreVars.v index 562b478..d158f05 100644 --- a/src/HaskCoreVars.v +++ b/src/HaskCoreVars.v @@ -6,6 +6,7 @@ Generalizable All Variables. Require Import Preamble. Require Import General. Require Import Coq.Strings.String. +Require Import HaskLiteralsAndTyCons. (* GHC uses a single type for expression variables, type variables, and coercion variables; this is that type *) Variable CoreVar : Type. Extract Inlined Constant CoreVar => "Var.Var". @@ -13,3 +14,18 @@ Variable coreVar_eq : forall (a b:CoreVar), sumbool (a=b) (not (a=b)). Variable coreVarToString : CoreVar -> string. Extract Inlined Constant coreVarToString => "outputableToString". Instance CoreVarEqDecidable : EqDecidable CoreVar := { eqd_dec := coreVar_eq }. Instance CoreVarToString : ToString CoreVar := { toString := coreVarToString }. + +Variable CoreTyCon : Type. Extract Inlined Constant CoreTyCon => "TyCon.TyCon". + +(* because Haskell's 3-tuples (triples) are distinct from both ((x,y),z) and (x,(y,z)), we need a new type: *) +Inductive triple {A B C:Type} := +| mkTriple : A -> B -> C -> triple. +Notation "a ** b ** c" := (mkTriple a b c) (at level 20). +Extract Inductive triple => "(,,)" [ "(,,)" ]. + +Inductive CoreAltCon := +| DataAlt : CoreDataCon -> CoreAltCon +| LitAlt : HaskLiteral -> CoreAltCon +| DEFAULT : CoreAltCon. +Extract Inductive CoreAltCon => + "CoreSyn.AltCon" [ "CoreSyn.DataAlt" "CoreSyn.LitAlt" "CoreSyn.DEFAULT" ].