2024-07-24 11:17:45 +02:00
|
|
|
|
(* Define the abstract syntax of the calculus
|
|
|
|
|
* by inductive definition of type-terms
|
|
|
|
|
* and expression-terms.
|
|
|
|
|
*)
|
|
|
|
|
|
|
|
|
|
From Coq Require Import Strings.String.
|
|
|
|
|
Module Terms.
|
|
|
|
|
|
|
|
|
|
(* types *)
|
2024-07-25 11:04:56 +02:00
|
|
|
|
Inductive type_term : Type :=
|
|
|
|
|
| type_id : string -> type_term
|
|
|
|
|
| type_var : string -> type_term
|
|
|
|
|
| type_fun : type_term -> type_term -> type_term
|
|
|
|
|
| type_univ : string -> type_term -> type_term
|
|
|
|
|
| type_spec : type_term -> type_term -> type_term
|
2024-07-25 12:41:44 +02:00
|
|
|
|
| type_morph : type_term -> type_term -> type_term
|
2024-07-25 11:04:56 +02:00
|
|
|
|
| type_ladder : type_term -> type_term -> type_term
|
2024-07-24 11:17:45 +02:00
|
|
|
|
.
|
|
|
|
|
|
|
|
|
|
(* expressions *)
|
2024-07-25 11:04:56 +02:00
|
|
|
|
Inductive expr_term : Type :=
|
|
|
|
|
| expr_var : string -> expr_term
|
|
|
|
|
| expr_ty_abs : string -> expr_term -> expr_term
|
|
|
|
|
| expr_ty_app : expr_term -> type_term -> expr_term
|
2024-08-22 01:04:40 +02:00
|
|
|
|
| expr_abs : string -> type_term -> expr_term -> expr_term
|
|
|
|
|
| expr_morph : string -> type_term -> expr_term -> expr_term
|
|
|
|
|
| expr_app : expr_term -> expr_term -> expr_term
|
2024-09-08 14:48:36 +02:00
|
|
|
|
| expr_let : string -> expr_term -> expr_term -> expr_term
|
2024-07-25 11:04:56 +02:00
|
|
|
|
| expr_ascend : type_term -> expr_term -> expr_term
|
|
|
|
|
| expr_descend : type_term -> expr_term -> expr_term
|
2024-07-24 11:17:45 +02:00
|
|
|
|
.
|
|
|
|
|
|
2024-09-16 15:14:53 +02:00
|
|
|
|
(* values *)
|
|
|
|
|
Inductive is_abs_value : expr_term -> Prop :=
|
|
|
|
|
| VAbs_Var : forall x,
|
|
|
|
|
(is_abs_value (expr_var x))
|
2024-09-04 12:47:03 +02:00
|
|
|
|
|
2024-09-16 15:14:53 +02:00
|
|
|
|
| VAbs_Abs : forall x τ e,
|
|
|
|
|
(is_abs_value (expr_abs x τ e))
|
2024-09-04 12:47:03 +02:00
|
|
|
|
|
2024-09-16 15:14:53 +02:00
|
|
|
|
| VAbs_Morph : forall x τ e,
|
|
|
|
|
(is_abs_value (expr_morph x τ e))
|
2024-09-04 12:47:03 +02:00
|
|
|
|
|
2024-09-16 15:14:53 +02:00
|
|
|
|
| VAbs_TypAbs : forall τ e,
|
|
|
|
|
(is_abs_value (expr_ty_abs τ e))
|
|
|
|
|
.
|
2024-09-04 12:47:03 +02:00
|
|
|
|
|
2024-08-18 10:27:21 +02:00
|
|
|
|
Inductive is_value : expr_term -> Prop :=
|
2024-09-16 15:14:53 +02:00
|
|
|
|
| V_Abs : forall e,
|
|
|
|
|
(is_abs_value e) ->
|
|
|
|
|
(is_value e)
|
2024-08-18 10:27:21 +02:00
|
|
|
|
|
|
|
|
|
| V_Ascend : forall τ e,
|
2024-09-16 15:14:53 +02:00
|
|
|
|
(is_abs_value e) ->
|
2024-08-18 10:27:21 +02:00
|
|
|
|
(is_value (expr_ascend τ e))
|
2024-07-24 11:17:45 +02:00
|
|
|
|
|
2024-09-16 15:14:53 +02:00
|
|
|
|
| V_Descend : forall τ e,
|
|
|
|
|
(is_abs_value e) ->
|
|
|
|
|
(is_value (expr_descend τ e))
|
|
|
|
|
.
|
2024-08-22 01:04:40 +02:00
|
|
|
|
|
2024-07-27 13:28:52 +02:00
|
|
|
|
Declare Scope ladder_type_scope.
|
|
|
|
|
Declare Scope ladder_expr_scope.
|
|
|
|
|
Declare Custom Entry ladder_type.
|
2024-08-18 10:27:21 +02:00
|
|
|
|
Declare Custom Entry ladder_expr.
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-08-22 09:57:47 +02:00
|
|
|
|
Notation "[< t >]" := t
|
2024-08-18 10:27:21 +02:00
|
|
|
|
(t custom ladder_type at level 80) : ladder_type_scope.
|
|
|
|
|
Notation "'∀' x ',' t" := (type_univ x t)
|
|
|
|
|
(t custom ladder_type at level 80, in custom ladder_type at level 80, x constr).
|
|
|
|
|
Notation "'<' σ τ '>'" := (type_spec σ τ)
|
|
|
|
|
(in custom ladder_type at level 80, left associativity) : ladder_type_scope.
|
|
|
|
|
Notation "'(' τ ')'" := τ
|
|
|
|
|
(in custom ladder_type at level 70) : ladder_type_scope.
|
|
|
|
|
Notation "σ '->' τ" := (type_fun σ τ)
|
|
|
|
|
(in custom ladder_type at level 75, right associativity) : ladder_type_scope.
|
|
|
|
|
Notation "σ '->morph' τ" := (type_morph σ τ)
|
|
|
|
|
(in custom ladder_type at level 75, right associativity, τ at level 80) : ladder_type_scope.
|
|
|
|
|
Notation "σ '~' τ" := (type_ladder σ τ)
|
|
|
|
|
(in custom ladder_type at level 70, right associativity) : ladder_type_scope.
|
|
|
|
|
Notation "'$' x '$'" := (type_id x%string)
|
|
|
|
|
(in custom ladder_type at level 0, x constr) : ladder_type_scope.
|
|
|
|
|
Notation "'%' x '%'" := (type_var x%string)
|
|
|
|
|
(in custom ladder_type at level 0, x constr) : ladder_type_scope.
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-08-22 09:57:47 +02:00
|
|
|
|
Notation "[{ e }]" := e
|
2024-08-18 10:27:21 +02:00
|
|
|
|
(e custom ladder_expr at level 80) : ladder_expr_scope.
|
|
|
|
|
Notation "'%' x '%'" := (expr_var x%string)
|
|
|
|
|
(in custom ladder_expr at level 0, x constr) : ladder_expr_scope.
|
|
|
|
|
Notation "'Λ' t '↦' e" := (expr_ty_abs t e)
|
2024-09-08 15:28:44 +02:00
|
|
|
|
(in custom ladder_expr at level 10, t constr, e custom ladder_expr at level 80) : ladder_expr_scope.
|
|
|
|
|
Notation "'λ' x τ '↦' e" := (expr_abs x τ e)
|
|
|
|
|
(in custom ladder_expr at level 10, x constr, τ custom ladder_type at level 99, e custom ladder_expr at level 99) :ladder_expr_scope.
|
|
|
|
|
Notation "'λ' x τ '↦morph' e" := (expr_morph x τ e)
|
|
|
|
|
(in custom ladder_expr at level 10, x constr, τ custom ladder_type at level 99, e custom ladder_expr at level 99) :ladder_expr_scope.
|
|
|
|
|
Notation "'let' x ':=' e 'in' t" := (expr_let x e t)
|
|
|
|
|
(in custom ladder_expr at level 20, x constr, e custom ladder_expr at level 99, t custom ladder_expr at level 99) : ladder_expr_scope.
|
|
|
|
|
Notation "e 'as' τ" := (expr_ascend τ e)
|
|
|
|
|
(in custom ladder_expr at level 30, e custom ladder_expr, τ custom ladder_type at level 99) : ladder_expr_scope.
|
2024-09-16 17:54:24 +02:00
|
|
|
|
Notation "e 'des' τ" := (expr_descend τ e)
|
|
|
|
|
(in custom ladder_expr at level 30, e custom ladder_expr, τ custom ladder_type at level 99) : ladder_expr_scope.
|
2024-09-08 15:28:44 +02:00
|
|
|
|
Notation "e1 e2" := (expr_app e1 e2)
|
|
|
|
|
(in custom ladder_expr at level 50) : ladder_expr_scope.
|
|
|
|
|
Notation "'(' e ')'" := e
|
|
|
|
|
(in custom ladder_expr at level 0) : ladder_expr_scope.
|
2024-08-22 01:04:40 +02:00
|
|
|
|
|
|
|
|
|
(* EXAMPLES *)
|
|
|
|
|
|
2024-07-27 13:28:52 +02:00
|
|
|
|
Open Scope ladder_type_scope.
|
2024-08-18 10:27:21 +02:00
|
|
|
|
Open Scope ladder_expr_scope.
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-08-22 09:57:47 +02:00
|
|
|
|
Check [< ∀"α", (< $"Seq"$ %"α"% > ~ < $"List"$ %"α"% >) >].
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-09-08 15:28:44 +02:00
|
|
|
|
Definition polymorphic_identity1 : expr_term := [{ Λ"T" ↦ λ"x"%"T"% ↦ (%"x"%) }].
|
2024-08-22 09:57:47 +02:00
|
|
|
|
Definition polymorphic_identity2 : expr_term := [{ Λ"T" ↦ λ"y"%"T"% ↦ %"y"% }].
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-08-18 10:27:21 +02:00
|
|
|
|
Compute polymorphic_identity1.
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-08-18 10:27:21 +02:00
|
|
|
|
Close Scope ladder_type_scope.
|
|
|
|
|
Close Scope ladder_expr_scope.
|
2024-07-27 13:28:52 +02:00
|
|
|
|
|
2024-07-24 11:17:45 +02:00
|
|
|
|
End Terms.
|