“Collapsing” state in a functional map?

喜你入骨 提交于 2019-12-10 18:28:23

问题


This question is inspired by Software Foundations, but isn't about an exercise. That I know about, so far.

The Imp chapter ("Simple Imperative Programs") uses functional maps (from the Maps chapter ("Total and Partial Maps")) for the environment of the simpl.[1] programming language.

Inductive id : Type :=
  | Id : string -> id.

Definition total_map (A:Type) := id -> A.

Definition t_empty {A:Type} (v : A) : total_map A :=
  (fun _ => v).

Definition t_update {A:Type} (m : total_map A) (x : id) (v : A) :=
  fun x' => if beq_id x x' then v else m x'.

Reasoning about programs in the Imp chapter requires manipulating program states of the form t_update (t_update (...) x 4) y 12), which becomes tedious because the state grows with every assignment, even if what the assignment is "implementing" is an update of an existing variable.

In theory, it should be possible to collapse program states down to single t_update's for each variable, and I've gone so far as to prove a couple of lemmas that go in the right direction.

Require Import Coq.Logic.FunctionalExtensionality.

Lemma swap_updates : forall A (m : total_map A) x vx y vy,
  x <> y -> t_update (t_update m x vx) y vy = t_update (t_update m y vy) x vx.
Proof.
  intros A m x vx y vy H. apply functional_extensionality. intros k.
  unfold t_update. destruct (beq_id y k) eqn:Eqyk.
  - destruct (beq_id x k) eqn:Eqxk.
    + apply beq_id_true_iff in Eqyk. apply beq_id_true_iff in Eqxk. subst.
      contradiction.
    + reflexivity.
  - destruct (beq_id x k) eqn:Eqxk; reflexivity.
Qed.

Lemma collapse_updates : forall A (m : total_map A) x v1 v2,
  t_update (t_update m x v1) x v2 = t_update m x v2.
Proof.
  intros A m x v1 v2. apply functional_extensionality. intros k.
  unfold t_update. destruct (beq_id x k); reflexivity.
Qed.

But I don't know how to state (much less prove or use, for that matter) a theorem that t_update (... (t_update (...) x 4) ...) x 27 is equal to t_update (...) x 27; i.e. with the useless bindings for x removed.

Any ideas?

[1] I just saw this. I'm leaving it, but I'm considering taking a vacation from Coq soon.


回答1:


But I don't know how to state (much less prove or use, for that matter) a theorem that t_update (... (t_update (...) x 4) ...) x 27 is equal to t_update (...) x 27; i.e. with the useless bindings for x removed.

First of all, it's not necessary to state a theorem like that, it's sufficient to be able to solve equations of the form

t_update ... (t_update ...) = t_update ... (t_update ...)

to achieve your goals. This is what does the following custom tactic.

Require Import Coq.Logic.FunctionalExtensionality.

Ltac t_solve :=
  intros; apply functional_extensionality; intros; unfold t_update;
  repeat lazymatch goal with
  |- context [beq_id ?v ?x] =>
    case_eq (beq_id v x);
    [rewrite beq_id_true_iff | rewrite beq_id_false_iff]; intros
  end;
  subst; (congruence || discriminate).

We know that states are just ladders of nested if beq_id ? ? then ... else ... expressions so if we destruct all the beq_id ? ? subexpressions using the context mechanism then some standard tactics are smart enough to finish the work for us. This solution won't work if the number on variables is large enough, but should be sufficient to work through Software Foundations.




回答2:


One way to do this is to proceed via an intermediate representation where we can more easily manipulate the updates. I did so by defining updates in terms of a list of (id, A) pairs (l_update), and then proved that we can add to the list of updates and replace any existing updates (add_update does this). The simplification is finally a tactic that first rewrites everything to use l_update, then computes the list of updates (which doesn't have duplicate IDs) by simplifying add_update, and then finally computes l_update down to t_updates (since it happens to be defined that way). In a real development it would be a bit easier since you'd be using strings for the IDs and the if string_dec "a" "b" then ... else ... calls should compute away, whereas I have to use a rewrite since the decision function is an axiom.

Require Import Coq.Logic.FunctionalExtensionality.
Require Import List.

Variable id:Type.
Axiom id_eq_dec : forall (x y:id), {x=y}+{x<>y}.

Variable A:Type.

Definition total_map := id -> A.

Definition t_update (m : total_map) (x : id) (v : A) : total_map :=
  fun x' => if id_eq_dec x x' then v else m x'.

Definition Updates := list (id*A).

Fixpoint l_update (m:total_map) (us:Updates) : total_map :=
  match us with
  | nil => m
  | (x,a)::us' => t_update (l_update m us') x a
  end.

Fixpoint add_update (us:Updates) x a : Updates :=
  match us with
  | nil => (x, a)::nil
  | (x',a')::us' => if id_eq_dec x x' then
                    add_update us' x a
                  else (x',a') :: add_update us' x a
  end.

Lemma t_update_neq : forall m x v x',
    x <> x' ->
    t_update m x v x' = m x'.
Proof.
  unfold t_update; intros.
  destruct (id_eq_dec x x'); congruence.
Qed.

Lemma t_update_eq : forall m x v,
    t_update m x v x = v.
Proof.
  unfold t_update; intros.
  destruct (id_eq_dec x x); congruence.
Qed.

Hint Rewrite t_update_neq using (solve [ auto ] ) : upd.
Hint Rewrite t_update_eq : upd.

Ltac cmp_id x x' := destruct (id_eq_dec x x'); subst; simpl;
                    try autorewrite with upd;
                    try congruence; auto.

Lemma l_update_add_other_update: forall us x a m a' x',
    x <> x' ->
    l_update m (add_update us x a) x' = l_update m (add_update us x a') x'.
Proof.
  induction us; simpl; intros.
  autorewrite with upd; auto.
  destruct a.
  cmp_id x i.
  cmp_id i x'.
Qed.

Hint Resolve l_update_add_other_update.

Lemma l_update_add_update : forall us m x a,
    l_update m (add_update us x a) x = a.
Proof.
  induction us; simpl; intros;
    autorewrite with upd;
    auto.
  destruct a; simpl.
  cmp_id x i.
Qed.

Hint Rewrite l_update_add_update : upd.

Lemma l_update_add_update_neq : forall us m x x' a,
    x <> x' ->
    l_update m (add_update us x a) x' = l_update m us x'.
Proof.
  induction us; intros.
  simpl; autorewrite with upd; auto.
  destruct a; simpl.
  cmp_id x i.
  cmp_id i x'.
Qed.

Hint Rewrite l_update_add_update_neq using solve [ auto ] : upd.

Theorem l_update_add : forall us m x a,
    t_update (l_update m us) x a =
    l_update m (add_update us x a).
Proof.
  induction us; intros; extensionality x'; simpl; auto.
  destruct a; simpl in *.
  rewrite IHus by auto.
  cmp_id x i.
  cmp_id i x'.
  cmp_id x x'.
  cmp_id i x'.
Qed.

Lemma if_id_eq_not : forall T (a b:T) i i',
    i <> i' ->
    (if id_eq_dec i i' then a else b) = b.
Proof.
  intros.
  destruct (id_eq_dec i i'); congruence.
Qed.

Lemma if_id_eq : forall T (a b:T) i,
    (if id_eq_dec i i then a else b) = a.
Proof.
  intros.
  destruct (id_eq_dec i i); congruence.
Qed.

Ltac simplify_updates m :=
  replace m with (l_update m nil) by auto;
  repeat rewrite l_update_add;
  repeat (simpl; rewrite ?if_id_eq_not, ?if_id_eq by auto).

Example update_chain1 : forall m x1 x2 x3 a1 a2 a3 a4 a5,
    x1 <> x2 ->
    x2 <> x3 ->
    x1 <> x3 ->
    t_update
      (t_update
        (t_update
            (t_update
              (t_update m x1 a1) x3 a2) x2 a3) x1 a4) x3 a5 =
    t_update
      (t_update
        (t_update m x3 a5) x1 a4) x2 a3.
Proof.
  intros.
  etransitivity.
  simplify_updates m.
  auto.
  auto.
Qed.



回答3:


How about something like this?

Definition almost_equal {A} x (m m' : total_map A) :=
  forall y, x <> y -> m y = m' y.

Lemma update_almost_equal : forall A (m : total_map A) x v1,
  almost_equal x (t_update m x v1) m.
Proof.
  intros A m x v1 y H.
  unfold t_update.
  destruct (beq_id x y) eqn:Eqxy.
  + rewrite beq_id_true_iff in Eqxy. congruence.
  + reflexivity.
Qed.

Lemma update_preserves_almost_equal : forall A (m m' : total_map A) x y v,
  almost_equal x m m' ->
  almost_equal x (t_update m y v) (t_update m' y v).
Proof.
  intros A m m' x y v H z H'.  
  unfold t_update.
  destruct (beq_id y z); auto.
Qed.

Lemma collapse_updates : forall x A (m m' : total_map A) v,
  almost_equal x m m' ->    
  t_update m x v = t_update m' x v.
Proof.
  intros x A m m' v Hae. apply functional_extensionality. intros k.
  unfold t_update.
  destruct (beq_id x k) eqn:Heq.
  + reflexivity.
  + rewrite beq_id_false_iff in Heq.
    apply Hae; exact Heq.
Qed.

You can rewrite with collapse_updates, and eauto can construct the proof of almost_equal.

Example ex1 :
  forall A x y z (v1 v2 v3 v4 : A) m,
      t_update (t_update (t_update (t_update m x v1) z v2) y v3) x v4
    = t_update (t_update (t_update m                 z v2) y v3) x v4.
Proof.
  intros.
  erewrite collapse_updates by eauto using update_almost_equal, update_preserves_almost_equal.
  reflexivity.
Qed.

Example ex2 :
  forall A x y z (v1 v2 v3 v4: A) m,
      t_update (t_update (t_update (t_update m y v1) z v2) y v3) x v4
    = t_update (t_update (t_update m                 z v2) y v3) x v4.
Proof.
  intros.
  erewrite (collapse_updates y) by eauto using update_almost_equal, update_preserves_almost_equal.
  reflexivity.
Qed.

A drawback is that you have to mention the variable you want to simplify on, e.g. collapse_updates y above. I guess this is because rewrite doesn't backtrack across different choices of subterms, so if you don't mention the variable it will try the outermost one and then give up if that fails. That meant that I couldn't figure out any way to use this as an autorewrite hint, although you can write an Ltac which tries all different variables:

Ltac simpl_updates :=
  repeat match goal with
  | [ |- context C1[t_update _ ?X _] ] =>
    erewrite (collapse_updates X) by eauto using update_almost_equal, update_preserves_almost_equal    
  end.

Example ex3 :
  forall A x y z w (v1 v2 v3 v4 v5 v6 v7 v8: A) m,
      t_update (t_update (t_update (t_update (t_update (t_update (t_update m x v1) y v2) z v3) x v4) y v5) w v6) y v7
    = t_update (t_update (t_update (t_update m z v3) x v4) w v6) y v7.
Proof.
  intros.
  simpl_updates.
  reflexivity.
Qed.


来源:https://stackoverflow.com/questions/47662000/collapsing-state-in-a-functional-map

标签
易学教程内所有资源均来自网络或用户发布的内容,如有违反法律规定的内容欢迎反馈
该文章没有解决你所遇到的问题?点击提问,说说你的问题,让更多的人一起探讨吧!