Module ars

This file collects definitions and theorems on abstract rewriting systems. These are particularly useful as we define the operational semantics as a small step semantics. This file defines a hint database ars containing some theorems on abstract rewriting systems.
Require Import Wf_nat.
Require Export tactics base.

Definitions

Section definitions.
  Context `(R : relation A).

An element is reducible if a step is possible.
  Definition red (x : A) := ∃ y, R x y.

An element is in normal form if no further steps are possible.
  Definition nf (x : A) := ¬red x.

The reflexive transitive closure.
  Inductive rtc : relation A :=
    | rtc_refl x : rtc x x
    | rtc_l x y z : R x yrtc y zrtc x z.

Reductions of exactly n steps.
  Inductive nsteps : natrelation A :=
    | nsteps_O x : nsteps 0 x x
    | nsteps_l n x y z : R x ynsteps n y znsteps (S n) x z.

Reduction of at most n steps.
  Inductive bsteps : natrelation A :=
    | bsteps_refl n x : bsteps n x x
    | bsteps_l n x y z : R x ybsteps n y zbsteps (S n) x z.

The transitive closure.
  Inductive tc : relation A :=
    | tc_once x y : R x ytc x y
    | tc_l x y z : R x ytc y ztc x z.

An element x is universally looping if all paths starting at x are infinite.
  CoInductive all_loop : AProp :=
    | all_loop_do_step x : red x → (∀ y, R x yall_loop y) → all_loop x.

An element x is existentally looping if some path starting at x is infinite.
  CoInductive ex_loop : AProp :=
    | ex_loop_do_step x y : R x yex_loop yex_loop x.
End definitions.

Hint Unfold nf red.

General theorems

Section rtc.
  Context `{R : relation A}.

  Hint Constructors rtc nsteps bsteps tc.

  Global Instance rtc_reflexive: Reflexive (rtc R).
  Proof. exact (@rtc_refl A R). Qed.
  Lemma rtc_transitive x y z : rtc R x yrtc R y zrtc R x z.
  Proof. induction 1; eauto. Qed.
  Global Instance: Transitive (rtc R).
  Proof. exact rtc_transitive. Qed.
  Lemma rtc_once x y : R x yrtc R x y.
  Proof. eauto. Qed.
  Lemma rtc_r x y z : rtc R x yR y zrtc R x z.
  Proof. intros. etransitivity; eauto. Qed.
  Lemma rtc_inv x z : rtc R x zx = z ∨ ∃ y, R x yrtc R y z.
  Proof. inversion_clear 1; eauto. Qed.
  Lemma rtc_ind_l (P : AProp) (z : A)
    (Prefl : P z) (Pstep : ∀ x y, R x yrtc R y zP yP x) :
    ∀ x, rtc R x zP x.
  Proof. induction 1; eauto. Qed.
  Lemma rtc_ind_r_weak (P : AAProp)
    (Prefl : ∀ x, P x x) (Pstep : ∀ x y z, rtc R x yR y zP x yP x z) :
    ∀ x z, rtc R x zP x z.
  Proof.
    cut (∀ y z, rtc R y z → ∀ x, rtc R x yP x yP x z).
    { eauto using rtc_refl. }
    induction 1; eauto using rtc_r.
  Qed.

  Lemma rtc_ind_r (P : AProp) (x : A)
    (Prefl : P x) (Pstep : ∀ y z, rtc R x yR y zP yP z) :
    ∀ z, rtc R x zP z.
  Proof.
    intros z p. revert x z p Prefl Pstep. refine (rtc_ind_r_weak _ _ _); eauto.
  Qed.

  Lemma rtc_inv_r x z : rtc R x zx = z ∨ ∃ y, rtc R x yR y z.
  Proof. revert z. apply rtc_ind_r; eauto. Qed.

  Lemma nsteps_once x y : R x ynsteps R 1 x y.
  Proof. eauto. Qed.
  Lemma nsteps_trans n m x y z :
    nsteps R n x ynsteps R m y znsteps R (n + m) x z.
  Proof. induction 1; simpl; eauto. Qed.
  Lemma nsteps_r n x y z : nsteps R n x yR y znsteps R (S n) x z.
  Proof. induction 1; eauto. Qed.
  Lemma nsteps_rtc n x y : nsteps R n x yrtc R x y.
  Proof. induction 1; eauto. Qed.
  Lemma rtc_nsteps x y : rtc R x y → ∃ n, nsteps R n x y.
  Proof. induction 1; firstorder eauto. Qed.

  Lemma bsteps_once n x y : R x ybsteps R (S n) x y.
  Proof. eauto. Qed.
  Lemma bsteps_plus_r n m x y :
    bsteps R n x ybsteps R (n + m) x y.
  Proof. induction 1; simpl; eauto. Qed.
  Lemma bsteps_weaken n m x y :
    nmbsteps R n x ybsteps R m x y.
  Proof.
    intros. rewrite (Minus.le_plus_minus n m); auto using bsteps_plus_r.
  Qed.

  Lemma bsteps_plus_l n m x y :
    bsteps R n x ybsteps R (m + n) x y.
  Proof. apply bsteps_weaken. auto with arith. Qed.
  Lemma bsteps_S n x y : bsteps R n x ybsteps R (S n) x y.
  Proof. apply bsteps_weaken. lia. Qed.
  Lemma bsteps_trans n m x y z :
    bsteps R n x ybsteps R m y zbsteps R (n + m) x z.
  Proof. induction 1; simpl; eauto using bsteps_plus_l. Qed.
  Lemma bsteps_r n x y z : bsteps R n x yR y zbsteps R (S n) x z.
  Proof. induction 1; eauto. Qed.
  Lemma bsteps_rtc n x y : bsteps R n x yrtc R x y.
  Proof. induction 1; eauto. Qed.
  Lemma rtc_bsteps x y : rtc R x y → ∃ n, bsteps R n x y.
  Proof. induction 1; [exists 0; constructor|]. naive_solver eauto. Qed.
  Lemma bsteps_ind_r (P : natAProp) (x : A)
    (Prefl : ∀ n, P n x)
    (Pstep : ∀ n y z, bsteps R n x yR y zP n yP (S n) z) :
    ∀ n z, bsteps R n x zP n z.
  Proof.
    cut (∀ m y z, bsteps R m y z → ∀ n,
      bsteps R n x y → (∀ m', nm' ∧ m' ≤ n + mP m' y) → P (n + m) z).
    { intros help ?. change n with (0 + n). eauto. }
    induction 1 as [|m x' y z p2 p3 IH]; [by eauto with arith|].
    intros n p1 H. rewrite <-plus_n_Sm.
    apply (IH (S n)); [by eauto using bsteps_r |].
    intros [|m'] [??]; [lia |]. apply Pstep with x'.
    * apply bsteps_weaken with n; intuition lia.
    * done.
    * apply H; intuition lia.
  Qed.


  Lemma tc_transitive x y z : tc R x ytc R y ztc R x z.
  Proof. induction 1; eauto. Qed.
  Global Instance: Transitive (tc R).
  Proof. exact tc_transitive. Qed.
  Lemma tc_r x y z : tc R x yR y ztc R x z.
  Proof. intros. etransitivity; eauto. Qed.
  Lemma tc_rtc_l x y z : rtc R x ytc R y ztc R x z.
  Proof. induction 1; eauto. Qed.
  Lemma tc_rtc_r x y z : tc R x yrtc R y ztc R x z.
  Proof. intros Hxy Hyz. revert x Hxy. induction Hyz; eauto using tc_r. Qed.
  Lemma tc_rtc x y : tc R x yrtc R x y.
  Proof. induction 1; eauto. Qed.

  Lemma all_loop_red x : all_loop R xred R x.
  Proof. destruct 1; auto. Qed.
  Lemma all_loop_step x y : all_loop R xR x yall_loop R y.
  Proof. destruct 1; auto. Qed.
  Lemma all_loop_rtc x y : all_loop R xrtc R x yall_loop R y.
  Proof. induction 2; eauto using all_loop_step. Qed.
  Lemma all_loop_alt x :
    all_loop R x ↔ ∀ y, rtc R x yred R y.
  Proof.
    split; [eauto using all_loop_red, all_loop_rtc|].
    intros H. cut (∀ z, rtc R x zall_loop R z); [eauto|].
    cofix FIX. constructor; eauto using rtc_r.
  Qed.

End rtc.

Hint Constructors rtc nsteps bsteps tc : ars.
Hint Resolve rtc_once rtc_r tc_r rtc_transitive tc_rtc_l tc_rtc_r
  tc_rtc bsteps_once bsteps_r bsteps_refl bsteps_trans : ars.

Theorems on sub relations

Section subrel.
  Context {A} (R1 R2 : relation A).
  Notation subrel := (∀ x y, R1 x yR2 x y).
  Lemma red_subrel x : subrelred R1 xred R2 x.
  Proof. intros ? [y ?]; eauto. Qed.
  Lemma nf_subrel x : subrelnf R2 xnf R1 x.
  Proof. intros ? H1 H2; destruct H1; by apply red_subrel. Qed.
End subrel.

Theorems on well founded relations

Notation wf := well_founded.

Section wf.
  Context `{R : relation A}.

A trick by Thomas Braibant to compute with well-founded recursions: it lazily adds 2^n Acc_intro constructors in front of a well foundedness proof, so that the actual proof is never reached in practise.
  Fixpoint wf_guard (n : nat) (wfR : wf R) : wf R :=
    match n with
    | 0 => wfR
    | S n => λ x, Acc_intro xy _, wf_guard n (wf_guard n wfR) y)
    end.

  Lemma wf_projected `(R2 : relation B) (f : AB) :
    (∀ x y, R x yR2 (f x) (f y)) →
    wf R2wf R.
  Proof.
    intros Hf Hwf.
    cut (∀ y, Acc R2 y → ∀ x, y = f xAcc R x).
    { intros aux x. apply (aux (f x)); auto. }
    induction 1 as [y _ IH]. intros x ?. subst.
    constructor. intros. apply (IH (f y)); auto.
  Qed.

End wf.

Strategy 100 [wf_guard].