(* Throughout the course, we will be using the [stdpp] library to provide some useful and common features. We will introduce you to its features as we need them. *) From stdpp Require Import base tactics numbers strings. From stdpp Require relations. From semantics.lib Require Import maps. (** * Exercise sheet 0 *) (* We are using Coq's notion of integers, [Z]. All the standard operations, like [+] and [*], are defined on it. *) Inductive expr := | Const (z : Z) | Plus (e1 e2 : expr) | Mul (e1 e2 : expr). (** Exercise 1: Arithmetics *) Fixpoint expr_eval (e : expr) : Z := (* TODO: write the function *) match e with | Const z => z | Plus a b => (expr_eval a) + (expr_eval b) | Mul a b => (expr_eval a) * (expr_eval b) end. (* Check expr_eval (Mul (Const 3) (Const 2)). *) (** Now let's define some notation to make it look nice! *) (* We declare a so-called notation scope, so that we can still use the nice notations for addition on natural numbers [nat] and integers [Z]. *) Declare Scope expr. Delimit Scope expr with E. Notation "e1 + e2" := (Plus e1%Z e2%Z) : expr. Notation "e1 * e2" := (Mul e1%Z e2%Z) : expr. (* We can use our nice notation to write expressions! (note the [%E] to tell Coq to parse this as an arithmetical expression with the notations we just defined). *) (* Check (Const 5 + Const 5)%E. *) (* The notation also still works for [Z] and [nat]: *) (* Check (5 + 5)%Z. *) (* Check (5 + 5). *) (* As an exercise, rephrase the following lemmas using the newly defined notation *) Lemma expr_eval_test: expr_eval (Plus (Const (-4)) (Const 5)) = 1%Z. Proof. (* Apply the definition of expr_eval to the terms of the expression *) simpl. (* Use the linear algebra solver to finish the proof *) lia. Qed. Lemma plus_eval_comm e1 e2 : expr_eval (Plus e1 e2) = expr_eval (Plus e2 e1). Proof. simpl. rewrite Z.add_comm. reflexivity. Qed. Lemma plus_syntax_not_comm : Plus (Const 0) (Const 1) ≠ Plus (Const 1) (Const 0). Proof. discriminate. Qed. (** Exercise 2: Open arithmetical expressions *) (* Finally, we introduce variables into our arithmetic expressions. Variables are of Coq's [string] type. *) Inductive expr' := | Var (x: string) | Const' (z : Z) | Plus' (e1 e2 : expr') | Mul' (e1 e2 : expr'). (* Note: I have the habit of first defining these in Prop, before defining a function in bool if needed. I guess this makes more sense, though, as you ensure it is decidable *) (* We call an expression closed under the list X, if it only contains variables in X *) Fixpoint is_closed (X: list string) (e: expr') : bool := match e with | Var x => bool_decide (x ∈ X) | Const' z => true | Plus' e1 e2 => is_closed X e1 && is_closed X e2 | Mul' e1 e2 => is_closed X e1 && is_closed X e2 end. Definition closed X e := is_closed X e = true. (* Some examples of closed terms. *) Lemma example_no_vars_closed: closed [] (Plus' (Const' 3) (Const' 5)). Proof. unfold closed. simpl. done. Qed. Lemma example_some_vars_closed: closed ["x"; "y"] (Plus' (Var "x") (Var "y")). Proof. unfold closed. simpl. done. Qed. Lemma example_not_closed: ¬ closed ["x"] (Plus' (Var "x") (Var "y")). Proof. unfold closed. simpl. done. Qed. Lemma closed_mono X Y e: X ⊆ Y → closed X e → closed Y e. Proof. unfold closed. intros Hsub. induction e as [ x | z | e1 IHe1 e2 IHe2 | e1 IHe1 e2 IHe2]. simpl. - rewrite (bool_decide_eq_true (elem_of x X)). rewrite (bool_decide_eq_true (elem_of x Y)). auto. - done. - simpl. intro H. rewrite andb_true_intro; auto. split. rewrite IHe1; auto. apply (andb_prop _ _ H). rewrite IHe2; auto. apply (andb_prop _ _ H). - simpl. intro H. rewrite andb_true_intro; auto. split. rewrite IHe1; auto. apply (andb_prop _ _ H). rewrite IHe2; auto. apply (andb_prop _ _ H). Qed. (* we define a substitution operation on open expressions *) Fixpoint subst (e: expr') (x: string) (e': expr') : expr' := match e with | Var y => if (bool_decide (x = y)) then e' else Var y | Const' z => Const' z | Plus' e1 e2 => Plus' (subst e1 x e') (subst e2 x e') | Mul' e1 e2 => Mul' (subst e1 x e') (subst e2 x e') end. Lemma closed_plus {e1 e2 X}: closed X (Plus' e1 e2) -> (closed X e1) ∧ (closed X e2). Proof. intro H_closed. (* Split H_closed into two hypotheses, one for `closed X e1` and one for `closewd X e2` *) unfold closed in H_closed. unfold is_closed in H_closed. rewrite andb_true_iff in H_closed. destruct H_closed as [H_closed_e1 H_closed_e2]. split. assumption. assumption. Qed. Lemma closed_mul {e1 e2 X}: closed X (Mul' e1 e2) -> (closed X e1) ∧ (closed X e2). Proof. intro H_closed. (* Split H_closed into two hypotheses, one for `closed X e1` and one for `closewd X e2` *) unfold closed in H_closed. unfold is_closed in H_closed. rewrite andb_true_iff in H_closed. destruct H_closed as [H_closed_e1 H_closed_e2]. split. assumption. assumption. Qed. Lemma subst_closed e e' x X: closed X e → ¬ (x ∈ X) → subst e x e' = e. Proof. intro H_closed. intro H_x_free. induction e as [ y | y | e1 IHe1 e2 IHe2 | e1 IHe1 e2 IHe2 ]. simpl. - rewrite <- decide_bool_decide. rewrite decide_False. auto. intro H_eq. (* Now we just need to prove a contradiction from x=y, x not in X, closed X (Var y) *) unfold closed in H_closed. unfold is_closed in H_closed. rewrite H_eq in H_x_free. rewrite (bool_decide_eq_true (y ∈ X)) in H_closed. (* Contradiction found: y ∈ X and y ∉ X; NOTE: y ∉ X is syntax sugar for (y ∈ X) -> False *) exact (H_x_free H_closed). - unfold subst; auto. - apply closed_plus in H_closed. destruct H_closed as [H_closed_e1 H_closed_e2]. simpl. rewrite IHe1. rewrite IHe2. auto. assumption. assumption. - apply closed_mul in H_closed. destruct H_closed as [H_closed_e1 H_closed_e2]. simpl. rewrite IHe1. rewrite IHe2. auto. assumption. assumption. Qed. (* To evaluate an arithmetic expression, we define an evaluation function [expr_eval], which maps them to integers. Since our expressions contain variables, we pass a finite map as the argument, which is used to look up variables. The type of finite maps that we use is called [gmap]. *) Fixpoint expr_eval' (m: gmap string Z) (e : expr') : Z := match e with | Var x => default 0%Z (m !! x) (* this is the lookup operation on gmaps *) | Const' z => z | Plus' e1 e2 => (expr_eval' m e1) + (expr_eval' m e2) | Mul' e1 e2 => (expr_eval' m e1) * (expr_eval' m e2) end. (* Prove the following lemma which explains how substitution interacts with evaluation *) Lemma eval_subst_extend (m: gmap string Z) e x e': expr_eval' m (subst e x e') = expr_eval' (<[x := expr_eval' m e']> m) e. Proof. induction e as [ y | y | e1 IHe1 e2 IHe2 | e1 IHe1 e2 IHe2 ]. - simpl. (* Use the decidability of bool to inspect both possibilities *) destruct (bool_dec (bool_decide (x=y)) true) as [H_eq | H_neq]. (* If x = y *) apply bool_decide_eq_true in H_eq. rewrite <- decide_bool_decide; rewrite decide_True. rewrite H_eq. rewrite lookup_insert. unfold default. simpl. reflexivity. assumption. (* If x ≠ y *) apply not_true_is_false in H_neq. apply bool_decide_eq_false in H_neq. rewrite <- decide_bool_decide; rewrite decide_False; simpl. rewrite lookup_insert_ne. reflexivity. assumption. assumption. - auto. - simpl. rewrite IHe1. rewrite IHe2. reflexivity. - simpl. rewrite IHe1. rewrite IHe2. reflexivity. Qed.