X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskKinds.v;h=2200f517165589a3c832dba14e2d1cf205aec3cf;hp=1452045e63f730a656e02d3f5be22758982d021e;hb=b4857a6f575dffd5c9c9d5decbc21ff63a338270;hpb=48bc98e014fd0c21ca75017bf689e8e6e80f54e3 diff --git a/src/HaskKinds.v b/src/HaskKinds.v index 1452045..2200f51 100644 --- a/src/HaskKinds.v +++ b/src/HaskKinds.v @@ -12,21 +12,21 @@ Variable Note : Type. Extract Inlined Con Variable natToString : nat -> string. Extract Inlined Constant natToString => "natToString". Instance NatToStringInstance : ToString nat := { toString := natToString }. -(* Figure 7: production κ, ι *) Inductive Kind : Type := - | KindType : Kind (* ★ - the kind of coercions and the kind of types inhabited by [boxed] values *) - | KindTypeFunction : Kind -> Kind -> Kind (* ⇛ - type-function-space; things of kind X⇛Y are NOT inhabited by expressions*) + | KindStar : Kind (* ★ - the kind of coercions and the kind of types inhabited by [boxed] values *) + | KindArrow : Kind -> Kind -> Kind (* ⇛ - type-function-space; things of kind X⇛Y are NOT inhabited by expressions*) | KindUnliftedType : Kind (* not in the paper - this is the kind of unboxed non-tuple types *) | KindUnboxedTuple : Kind (* not in the paper - this is the kind of unboxed tuples *) - | KindArgType : Kind (* not in the paper - this is the lub of KindType and KindUnliftedType *) - | KindOpenType : Kind (* not in the paper - kind of all types (lifted, boxed, whatever) *). + | KindArgType : Kind (* not in the paper - this is the lub of KindStar and KindUnliftedType *) + | KindOpenType : Kind (* not in the paper - kind of all types (lifted, boxed, whatever) *) + . Open Scope string_scope. Fixpoint kindToString (k:Kind) : string := match k with - | KindType => "*" - | KindTypeFunction KindType k2 => "*=>"+++kindToString k2 - | KindTypeFunction k1 k2 => "("+++kindToString k1+++")=>"+++kindToString k2 + | KindStar => "*" + | KindArrow KindStar k2 => "*=>"+++kindToString k2 + | KindArrow k1 k2 => "("+++kindToString k1+++")=>"+++kindToString k2 | KindUnliftedType => "#" | KindUnboxedTuple => "(#)" | KindArgType => "??" @@ -34,8 +34,24 @@ Fixpoint kindToString (k:Kind) : string := end. Instance KindToString : ToString Kind := { toString := kindToString }. -Notation "'★'" := KindType. -Notation "a ⇛ b" := (KindTypeFunction a b). +Notation "'★'" := KindStar. +Notation "a ⇛ b" := (KindArrow a b). + +(* the kind of environment classifiers *) +Definition ECKind := ★ . +Opaque ECKind. + +Fixpoint kindToLatexMath (k:Kind) : LatexMath := + match k with + | ★ => rawLatexMath "\star" + | ★ ⇛ k2 => (rawLatexMath "\star\Rightarrow ")+++kindToLatexMath k2 + | k1 ⇛ k2 => (rawLatexMath "(")+++kindToLatexMath k1+++(rawLatexMath ")\Rightarrow ")+++kindToLatexMath k2 + | KindUnliftedType => rawLatexMath "\text{\tt{\#}}" + | KindUnboxedTuple => rawLatexMath "\text{\tt{(\#)}}" + | KindArgType => rawLatexMath "\text{\tt{??}}" + | KindOpenType => rawLatexMath "\text{\tt{?}}" + end. +Instance KindToLatexMath : ToLatexMath Kind := { toLatexMath := kindToLatexMath }. Instance KindEqDecidable : EqDecidable Kind. apply Build_EqDecidable. @@ -50,3 +66,10 @@ Instance KindEqDecidable : EqDecidable Kind. destruct v2; try (right; intro q; inversion q; fail) ; left ; auto. Defined. +(* splits a kind into a list of arguments and a result *) +Fixpoint splitKind (k:Kind) : ((list Kind) * Kind) := + match k with + | a ⇛ b => let (args,res) := splitKind b in ((a::args),res) + | _ => (nil, k) + end. +