Library Coq.ZArith.Zwf


Require Import ZArith_base.
Require Export Wf_nat.
Require Import Lia.
Local Open Scope Z_scope.

Well-founded relations on Z.
We define the following family of relations on Z x Z:
x (Zwf c) y iff x < y & c <= y

Definition Zwf (c x y:Z) := c <= y /\ x < y.

and we prove that (Zwf c) is well founded

Section wf_proof.

  Variable c : Z.

The proof of well-foundness is classic: we do the proof by induction on a measure in nat, which is here |x-c|

  Let f (z:Z) := Z.abs_nat (z - c).

  Lemma Zwf_well_founded : well_founded (Zwf c).
n= 0
inductive case

End wf_proof.

Hint Resolve Zwf_well_founded: datatypes.

We also define the other family of relations:
x (Zwf_up c) y iff y < x <= c

Definition Zwf_up (c x y:Z) := y < x <= c.

and we prove that (Zwf_up c) is well founded

Section wf_proof_up.

  Variable c : Z.

The proof of well-foundness is classic: we do the proof by induction on a measure in nat, which is here |c-x|

  Let f (z:Z) := Z.abs_nat (c - z).

  Lemma Zwf_up_well_founded : well_founded (Zwf_up c).

End wf_proof_up.

Hint Resolve Zwf_up_well_founded: datatypes.