ladder-calculus/coq/terms.v

71 lines
2.9 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.

(* 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 *)
Inductive type_term : Type :=
| type_unit : type_term
| type_id : string -> type_term
| type_var : string -> type_term
| type_num : nat -> 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
| type_morph : type_term -> type_term -> type_term
| type_ladder : type_term -> type_term -> type_term
.
(* expressions *)
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
| expr_tm_abs : string -> type_term -> expr_term -> expr_term
| expr_tm_abs_morph : string -> type_term -> expr_term -> expr_term
| expr_tm_app : expr_term -> expr_term -> expr_term
| expr_let : string -> type_term -> expr_term -> expr_term -> expr_term
| expr_ascend : type_term -> expr_term -> expr_term
| expr_descend : type_term -> expr_term -> expr_term
.
Coercion type_var : string >-> type_term.
Coercion expr_var : string >-> expr_term.
(*
Coercion type_var : string >-> type_term.
Coercion expr_var : string >-> expr_term.
*)
Declare Scope ladder_type_scope.
Declare Scope ladder_expr_scope.
Declare Custom Entry ladder_type.
Notation "[ e ]" := e (e custom ladder_type at level 80) : ladder_type_scope.
(* TODO: allow any variable names in notation, not just α,β,γ *)
Notation "'∀α.' τ" := (type_univ "α" τ) (in custom ladder_type at level 80) : ladder_type_scope.
Notation "'∀β.' τ" := (type_univ "β" τ) (in custom ladder_type at level 80) : ladder_type_scope.
Notation "'∀γ.' τ" := (type_univ "γ" τ) (in custom ladder_type at level 80) : ladder_type_scope.
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) : ladder_type_scope.
Notation "σ '~' τ" := (type_ladder σ τ) (in custom ladder_type at level 70, right associativity) : ladder_type_scope.
Notation "'α'" := (type_var "α") (in custom ladder_type at level 60, right associativity) : ladder_type_scope.
Notation "'β'" := (type_var "β") (in custom ladder_type at level 60, right associativity) : ladder_type_scope.
Notation "'γ'" := (type_var "γ") (in custom ladder_type at level 60, right associativity) : ladder_type_scope.
Open Scope ladder_type_scope.
Definition t1 : type_term := [ α.β.(α~β~γ)->β->(α->α)->β ].
Compute t1.
Close Scope ladder_type_scope.
End Terms.