From Coq Require Import Strings.String. Require Import terms. Require Import subst. Require Import typing. Include Terms. Include Subst. Include Typing. Module Smallstep. Reserved Notation " s '-->α' t " (at level 40). Reserved Notation " s '-->β' t " (at level 40). Inductive expr_alpha : expr_term -> expr_term -> Prop := | EAlpha_Rename : forall x x' τ e, (expr_abs x τ e) -->α (expr_abs x' τ (expr_subst x (expr_var x') e)) | EAlpha_TyRename : forall α α' e, (expr_ty_abs α e) -->α (expr_ty_abs α' (expr_specialize α (type_var α') e)) | EAlpha_SubAbs : forall x τ e e', (e -->α e') -> (expr_abs x τ e) -->α (expr_abs x τ e') | EAlpha_SubTyAbs : forall α e e', (e -->α e') -> (expr_ty_abs α e) -->α (expr_ty_abs α e') | EAlpha_SubApp1 : forall e1 e1' e2, (e1 -->α e1') -> (expr_app e1 e2) -->α (expr_app e1' e2) | EAlpha_SubApp2 : forall e1 e2 e2', (e2 -->α e2') -> (expr_app e1 e2) -->α (expr_app e1 e2') where "s '-->α' t" := (expr_alpha s t). Example a1 : polymorphic_identity1 -->α polymorphic_identity2. Proof. unfold polymorphic_identity1. unfold polymorphic_identity2. apply EAlpha_SubTyAbs. apply EAlpha_Rename. Qed. Inductive beta_step : expr_term -> expr_term -> Prop := | E_App1 : forall e1 e1' e2, e1 -->β e1' -> (expr_app e1 e2) -->β (expr_app e1' e2) | E_App2 : forall v1 e2 e2', (is_value v1) -> e2 -->β e2' -> (expr_app v1 e2) -->β (expr_app v1 e2') | E_TypApp : forall e e' τ, e -->β e' -> (expr_ty_app e τ) -->β (expr_ty_app e' τ) | E_TypAppLam : forall x e a, (expr_ty_app (expr_ty_abs x e) a) -->β (expr_specialize x a e) | E_AppLam : forall x τ e a, (expr_tm_app (expr_tm_abs x τ e) a) -->β (expr_subst x a e) | E_AppLet : forall x t e a, (expr_let x t a e) -->β (expr_subst x a e) where "s '-->β' t" := (beta_step s t). Inductive delta_step : expr_term -> expr_term -> Prop := | E_ImplicitCast : forall (Γ:context) (f:expr_term) (h:string) (a:expr_term) (τ:type_term) (s:type_term) (p:type_term), (context_contains Γ h (type_morph p s)) -> Γ |- f \is (type_fun s τ) -> Γ |- a \is p -> (expr_tm_app f a) -->δ (expr_tm_app f (expr_tm_app (expr_var h) a)) where "s '-->δ' t" := (delta_step s t). Inductive eval_step : expr_term -> expr_term -> Prop := | E_Beta : forall s t, (s -->β t) -> (s -->eval t) | E_Delta : forall s t, (s -->δ t) -> (s -->eval t) where "s '-->eval' t" := (eval_step s t). Inductive multi {X : Type} (R : X -> X -> Prop) : X -> X -> Prop := | Multi_Refl : forall (x : X), multi R x x | Multi_Step : forall (x y z : X), R x y -> multi R y z -> multi R x z. Notation " s -->α* t " := (multi expr_alpha s t) (at level 40). Notation " s -->β* t " := (multi beta_step s t) (at level 40). End Smallstep.