Defining recursive functions by Noetherian induction. This is a simplified
interface to the Wf module of Coq's standard library, where the functions
to be defined have non-dependent types, and function extensionality is assumed.
Require Import Axioms.
Require Import Init.Wf.
Require Import Wf_nat.
Set Implicit Arguments.
Section FIX.
Variables A B:
Type.
Variable R:
A ->
A ->
Prop.
Hypothesis Rwf:
well_founded R.
Variable F:
forall (
x:
A), (
forall (
y:
A),
R y x ->
B) ->
B.
Definition Fix (
x:
A) :
B :=
Wf.Fix Rwf (
fun (
x:
A) =>
B)
F x.
Theorem unroll_Fix:
forall x,
Fix x =
F (
fun (
y:
A) (
P:
R y x) =>
Fix y).
Proof.
End FIX.
Same, with a nonnegative measure instead of a well-founded ordering
Section FIXM.
Variables A B:
Type.
Variable measure:
A ->
nat.
Variable F:
forall (
x:
A), (
forall (
y:
A),
measure y <
measure x ->
B) ->
B.
Definition Fixm (
x:
A) :
B :=
Wf.Fix (
well_founded_ltof A measure) (
fun (
x:
A) =>
B)
F x.
Theorem unroll_Fixm:
forall x,
Fixm x =
F (
fun (
y:
A) (
P:
measure y <
measure x) =>
Fixm y).
Proof.
End FIXM.