ladder-calculus/coq/bbencode.v

83 lines
2.7 KiB
Coq
Raw Normal View History

From Coq Require Import Strings.String.
Require Import terms.
Require Import subst.
Require Import smallstep.
Include Terms.
Include Subst.
Include Smallstep.
2024-07-24 11:20:43 +02:00
(* let bb_zero = Λα ↦ λs: α->α ↦ λz: α ↦ z
* α.(α->α)->α->α
*)
Definition bb_zero : expr_term :=
2024-07-24 11:20:43 +02:00
(expr_ty_abs "α"
(expr_tm_abs "s" (type_fun (type_var "α") (type_var "α"))
(expr_tm_abs "z" (type_var "α")
(expr_var "z")))).
(* let bb_one = Λα ↦ λs: α->α ↦ λz: α ↦ s z
*)
Definition bb_one : expr_term :=
2024-07-24 11:20:43 +02:00
(expr_ty_abs "α"
(expr_tm_abs "s" (type_fun (type_var "α") (type_var "α"))
(expr_tm_abs "z" (type_var "α")
(expr_tm_app (expr_var "s") (expr_var "z"))))).
(* let bb_two = Λα ↦ λs: α->α ↦ λz: α ↦ s (s z)
*)
Definition bb_two : expr_term :=
2024-07-24 11:20:43 +02:00
(expr_ty_abs "α"
(expr_tm_abs "s" (type_fun (type_var "α") (type_var "α"))
(expr_tm_abs "z" (type_var "α")
(expr_tm_app (expr_var "s") (expr_tm_app (expr_var "s") (expr_var "z")))))).
Definition bb_succ : expr_term :=
(expr_tm_abs "n" (type_ladder (type_id "")
(type_ladder (type_id "BBNat")
(type_univ "α"
2024-07-24 11:20:43 +02:00
(type_fun (type_fun (type_var "α") (type_var "α"))
(type_fun (type_var "α") (type_var "α"))))))
(expr_ascend (type_ladder (type_id "") (type_id "BBNat"))
2024-07-24 11:20:43 +02:00
(expr_ty_abs "α"
(expr_tm_abs "s" (type_fun (type_var "α") (type_var "α"))
(expr_tm_abs "z" (type_var "α")
(expr_tm_app (expr_var "s")
(expr_tm_app (expr_tm_app
(expr_ty_app (expr_var "n") (type_var "α"))
(expr_var "s"))
(expr_var "z")))))))).
Definition e1 : expr_term :=
(expr_let "bb-zero" (type_ladder (type_id "")
(type_ladder (type_id "BBNat")
(type_univ "α"
2024-07-24 11:20:43 +02:00
(type_fun (type_fun (type_var "α") (type_var "α"))
(type_fun (type_var "α") (type_var "α"))))))
bb_zero
(expr_tm_app (expr_tm_app (expr_var "+") (expr_var "bb-zero")) (expr_var "bb-zero"))
).
Definition t1 : expr_term := (expr_tm_app (expr_var "x") (expr_var "x")).
2024-07-24 11:20:43 +02:00
Compute (expr_subst "x"
(expr_ty_abs "α" (expr_tm_abs "a" (type_var "α") (expr_var "a")))
bb_one
).
Example example_let_reduction :
e1 -->β (expr_tm_app (expr_tm_app (expr_var "+") bb_zero) bb_zero).
Proof.
apply E_AppLet.
Qed.
Compute (expr_tm_app bb_succ bb_zero) -->β bb_one.
Example example_succ :
(expr_tm_app bb_succ bb_zero) -->β bb_one.
Proof.
Admitted.