ladder-calculus/coq/smallstep.v
2024-08-21 20:19:03 +02:00

88 lines
2.4 KiB
Coq
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

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).
Reserved Notation " s '-->δ' t " (at level 40).
Reserved Notation " s '-->eval' t " (at level 40).
Inductive alpha_step : expr_term -> expr_term -> Prop :=
| E_Rename : forall x x' e,
(expr_tm_abs x e) -->α (expr_tm_abs x' (expr_subst x (type_var x')))
where "s '-->α' t" := (alpha_step s t).
Example a1 : polymorphic_identity1 -->α polymorphic_identity2.
Proof.
Qed.
Inductive beta_step : expr_term -> expr_term -> Prop :=
| E_App1 : forall e1 e1' e2,
e1 -->β e1' ->
(expr_tm_app e1 e2) -->β (expr_tm_app e1' e2)
| E_App2 : forall e1 e2 e2',
e2 -->β e2' ->
(expr_tm_app e1 e2) -->β (expr_tm_app e1 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 beta_step s t) (at level 40).
Notation " s -->δ* t " := (multi delta_step s t) (at level 40).
Notation " s -->eval* t " := (multi eval_step s t) (at level 40).
End Smallstep.