(* * This module defines the subtype relationship * * We distinguish between *representational* subtypes, * where any high-level type is a subtype of its underlying * representation type and *convertible* subtypes that * are compatible at high level, but have a different representation * that requires a conversion. *) Require Import debruijn. Require Import equiv. (** Subtyping *) Create HintDb subtype_hints. Reserved Notation "s ':<=' t" (at level 50). Reserved Notation "s '~<=' t" (at level 50). Local Open Scope ladder_type_scope. (* Representational Subtype *) Inductive is_repr_subtype : type_DeBruijn -> type_DeBruijn -> Prop := | TSubRepr_Refl : forall t t', (t === t') -> (t :<= t') | TSubRepr_Trans : forall x y z, (x :<= y) -> (y :<= z) -> (x :<= z) | TSubRepr_Ladder : forall x' x y, (x :<= y) -> ([< x' ~ x >] :<= y) where "s ':<=' t" := (is_repr_subtype s t). (* Convertible Subtype *) Inductive is_conv_subtype : type_DeBruijn -> type_DeBruijn -> Prop := | TSubConv_Refl : forall t t', (t === t') -> (t ~<= t') | TSubConv_Trans : forall x y z, (x ~<= y) -> (y ~<= z) -> (x ~<= z) | TSubConv_Ladder : forall x' x y, (x ~<= y) -> ([< x' ~ x >] ~<= y) | TSubConv_Morph : forall x y y', [< x ~ y >] ~<= [< x ~ y' >] where "s '~<=' t" := (is_conv_subtype s t). Hint Constructors is_repr_subtype :subtype_hints. Hint Constructors is_conv_subtype :subtype_hints. Hint Constructors type_eq :subtype_hints. (* Every Representational Subtype is a Convertible Subtype *) Lemma syn_sub_is_sem_sub : forall x y, (x :<= y) -> (x ~<= y). Proof. intros. induction H. all: eauto with subtype_hints. Qed.