diff options
| author | Rutger Broekhoff | 2024-06-26 20:50:18 +0200 |
|---|---|---|
| committer | Rutger Broekhoff | 2024-06-26 20:50:18 +0200 |
| commit | 73df1945b31c0beee88cf4476df4ccd09d31403b (patch) | |
| tree | ed00db26b711442e643f38b66888a3df56e33ebd /relations.v | |
| download | mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.tar.gz mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.zip | |
Import Coq project
Diffstat (limited to 'relations.v')
| -rw-r--r-- | relations.v | 73 |
1 files changed, 73 insertions, 0 deletions
diff --git a/relations.v b/relations.v new file mode 100644 index 0000000..9bfb7e6 --- /dev/null +++ b/relations.v | |||
| @@ -0,0 +1,73 @@ | |||
| 1 | From stdpp Require Import relations. | ||
| 2 | |||
| 3 | Definition deterministic {A} (R : relation A) := | ||
| 4 | ∀ x y1 y2, R x y1 → R x y2 → y1 = y2. | ||
| 5 | |||
| 6 | (* As in Baader and Nipkow (1998): *) | ||
| 7 | (* y is a normal form of x if x -->* y and y is in normal form. *) | ||
| 8 | Definition is_nf_of `(R : relation A) x y := (rtc R x y) ∧ nf R y. | ||
| 9 | |||
| 10 | (** The reflexive closure. *) | ||
| 11 | Inductive rc {A} (R : relation A) : relation A := | ||
| 12 | | rc_refl x : rc R x x | ||
| 13 | | rc_once x y : R x y → rc R x y. | ||
| 14 | |||
| 15 | Hint Constructors rc : core. | ||
| 16 | |||
| 17 | (* Baader and Nipkow, Definition 2.7.3. *) | ||
| 18 | Definition strongly_confluent {A} (R : relation A) := | ||
| 19 | ∀ x y1 y2, R x y1 → R x y2 → ∃ z, rtc R y1 z ∧ rc R y2 z. | ||
| 20 | |||
| 21 | (* Baader and Nipkow, Definition 2.1.4. *) | ||
| 22 | Definition semi_confluent {A} (R : relation A) := | ||
| 23 | ∀ x y1 y2, R x y1 → rtc R x y2 → ∃ z, rtc R y1 z ∧ rtc R y2 z. | ||
| 24 | |||
| 25 | (* Lemma 2.7.4 from Baader and Nipkow *) | ||
| 26 | Lemma strong__semi_confluence {A} (R : relation A) : | ||
| 27 | strongly_confluent R → semi_confluent R. | ||
| 28 | Proof. | ||
| 29 | intros Hstrc. | ||
| 30 | unfold strongly_confluent in Hstrc. | ||
| 31 | unfold semi_confluent. | ||
| 32 | intros x1 y1 xn Hxy1 [steps Hsteps] % rtc_nsteps. | ||
| 33 | revert x1 y1 xn Hxy1 Hsteps. | ||
| 34 | induction steps; intros x1 y1 xn Hxy1 Hsteps. | ||
| 35 | - inv Hsteps. exists y1. | ||
| 36 | split. | ||
| 37 | + apply rtc_refl. | ||
| 38 | + by apply rtc_once. | ||
| 39 | - inv Hsteps as [|? ? x2 ? Hx1x2 Hsteps']. | ||
| 40 | destruct (Hstrc x1 y1 x2 Hxy1 Hx1x2) as [y2 [Hy21 Hy22]]. | ||
| 41 | inv Hy22. | ||
| 42 | + exists xn. split; [|done]. | ||
| 43 | apply rtc_transitive with (y := y2); [done|]. | ||
| 44 | rewrite rtc_nsteps. by exists steps. | ||
| 45 | + destruct (IHsteps x2 y2 xn H Hsteps') as [yn [Hy2yn Hxnyn]]. | ||
| 46 | exists yn. split; [|done]. | ||
| 47 | by apply rtc_transitive with (y := y2). | ||
| 48 | Qed. | ||
| 49 | |||
| 50 | (* Copied from stdpp, originates from Baader and Nipkow (1998) *) | ||
| 51 | Definition cr {A} (R : relation A) := | ||
| 52 | ∀ x y, rtsc R x y → ∃ z, rtc R x z ∧ rtc R y z. | ||
| 53 | |||
| 54 | (* Part of Lemma 2.1.5 from Baader and Nipkow (1998) *) | ||
| 55 | Lemma semi_confluence__cr {A} (R : relation A) : | ||
| 56 | semi_confluent R → cr R. | ||
| 57 | Proof. | ||
| 58 | intros Hsc. | ||
| 59 | unfold semi_confluent in Hsc. | ||
| 60 | unfold cr. | ||
| 61 | intros x y [steps Hsteps] % rtc_nsteps. | ||
| 62 | revert x y Hsteps. | ||
| 63 | induction steps; intros x y Hsteps. | ||
| 64 | - inv Hsteps. by exists y. | ||
| 65 | - inv Hsteps. rename x into y', y into x, y0 into y. | ||
| 66 | apply IHsteps in H1 as H1'. destruct H1' as [z [Hz1 Hz2]]. | ||
| 67 | destruct H0. | ||
| 68 | + exists z. split; [|done]. | ||
| 69 | by apply rtc_l with (y := y). | ||
| 70 | + destruct (Hsc y y' z H Hz1) as [z' [Hz'1 Hz'2]]. | ||
| 71 | exists z'. split; [done|]. | ||
| 72 | by apply rtc_transitive with (y := z). | ||
| 73 | Qed. | ||