(* week-13_a-continuation-based-interpreter.v *) (* LPP 2024 - CS3234 2023-2024, Sem2 *) (* Olivier Danvy *) (* Version of 12 Apr 2024 *) (* ********** *) Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. Require Import Arith Bool List String Ascii. (* ***** *) Check (1 =? 2). Check Nat.eqb. Check (Nat.eqb 1 2). Check (1 <=? 2). Check Nat.leb. Check (Nat.leb 1 2). Check (1 arithmetic_expression | Plus : arithmetic_expression -> arithmetic_expression -> arithmetic_expression | Minus : arithmetic_expression -> arithmetic_expression -> arithmetic_expression. (* Source programs: *) Inductive source_program : Type := Source_program : arithmetic_expression -> source_program. (* ********** *) (* Semantics: *) Inductive expressible_value : Type := Expressible_nat : nat -> expressible_value | Expressible_msg : string -> expressible_value. (* ********** *) Definition specification_of_evaluate (evaluate : arithmetic_expression -> expressible_value) := (forall n : nat, evaluate (Literal n) = Expressible_nat n) /\ ((forall (ae1 ae2 : arithmetic_expression) (s1 : string), evaluate ae1 = Expressible_msg s1 -> evaluate (Plus ae1 ae2) = Expressible_msg s1) /\ (forall (ae1 ae2 : arithmetic_expression) (n1 : nat) (s2 : string), evaluate ae1 = Expressible_nat n1 -> evaluate ae2 = Expressible_msg s2 -> evaluate (Plus ae1 ae2) = Expressible_msg s2) /\ (forall (ae1 ae2 : arithmetic_expression) (n1 n2 : nat), evaluate ae1 = Expressible_nat n1 -> evaluate ae2 = Expressible_nat n2 -> evaluate (Plus ae1 ae2) = Expressible_nat (n1 + n2))) /\ ((forall (ae1 ae2 : arithmetic_expression) (s1 : string), evaluate ae1 = Expressible_msg s1 -> evaluate (Minus ae1 ae2) = Expressible_msg s1) /\ (forall (ae1 ae2 : arithmetic_expression) (n1 : nat) (s2 : string), evaluate ae1 = Expressible_nat n1 -> evaluate ae2 = Expressible_msg s2 -> evaluate (Minus ae1 ae2) = Expressible_msg s2) /\ (forall (ae1 ae2 : arithmetic_expression) (n1 n2 : nat), evaluate ae1 = Expressible_nat n1 -> evaluate ae2 = Expressible_nat n2 -> n1 evaluate (Minus ae1 ae2) = Expressible_msg (String.append "numerical underflow: -" (string_of_nat (n2 - n1)))) /\ (forall (ae1 ae2 : arithmetic_expression) (n1 n2 : nat), evaluate ae1 = Expressible_nat n1 -> evaluate ae2 = Expressible_nat n2 -> n1 evaluate (Minus ae1 ae2) = Expressible_nat (n1 - n2))). Definition specification_of_interpret (interpret : source_program -> expressible_value) := forall evaluate : arithmetic_expression -> expressible_value, specification_of_evaluate evaluate -> forall ae : arithmetic_expression, interpret (Source_program ae) = evaluate ae. (* ********** *) (* An evaluation function in delimited continuation-passing style: *) Fixpoint evaluate_cb (ae : arithmetic_expression) (k : nat -> expressible_value) : expressible_value := Expressible_msg "not implemented yet". (* Lemma fold_unfold_evaluate_cb_Literal : forall (n : nat) (k : nat -> expressible_value), evaluate_cb (Literal n) k = ... Proof. fold_unfold_tactic evaluate_cb. Qed. Lemma fold_unfold_evaluate_cb_Plus : forall (ae1 ae2 : arithmetic_expression) (k : nat -> expressible_value), evaluate_cb (Plus ae1 ae2) k = ... Proof. fold_unfold_tactic evaluate_cb. Qed. Lemma fold_unfold_evaluate_cb_Minus : forall (ae1 ae2 : arithmetic_expression) (k : nat -> expressible_value), evaluate_cb (Minus ae1 ae2) k = ... Proof. fold_unfold_tactic evaluate_cb. Qed. *) (* ***** *) Definition interpret (sp : source_program) : expressible_value := Expressible_msg "not implemented yet". (* ***** *) Lemma about_evaluate_cb : forall evaluate : arithmetic_expression -> expressible_value, specification_of_evaluate evaluate -> forall ae : arithmetic_expression, (exists n : nat, evaluate ae = Expressible_nat n /\ forall k : nat -> expressible_value, evaluate_cb ae k = k n) \/ (exists s : string, evaluate ae = Expressible_msg s /\ forall k : nat -> expressible_value, evaluate_cb ae k = Expressible_msg s). Proof. Abort. (* ***** *) Theorem interpret_satisfies_specification_of_interpret : specification_of_interpret interpret. Proof. Abort. (* ********** *) (* end of week-13_a-continuation-based-interpreter.v *)