blob: 4fb640827b6f09e7aa344151bc78851a19f397e0 (
plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
|
(***********************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
(* \VV/ *************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
(*i $Id$ i*)
(** This module proves the validity of
- well-founded recursion (also called course of values)
- well-founded induction
from a well-founded ordering on a given set *)
Require Logic.
Require LogicSyntax.
(** Well-founded induction principle on Prop *)
Chapter Well_founded.
Variable A : Set.
Variable R : A -> A -> Prop.
(** The accessibility predicate is defined to be non-informative *)
Inductive Acc : A -> Prop
:= Acc_intro : (x:A)((y:A)(R y x)->(Acc y))->(Acc x).
Lemma Acc_inv : (x:A)(Acc x) -> (y:A)(R y x) -> (Acc y).
NewDestruct 1; Trivial.
Defined.
(** the informative elimination :
[let Acc_rec F = let rec wf x = F x wf in wf] *)
Section AccRecType.
Variable P : A -> Type.
Variable F : (x:A)((y:A)(R y x)->(Acc y))->((y:A)(R y x)->(P y))->(P x).
Fixpoint Acc_rect [x:A;a:(Acc x)] : (P x)
:= (F x (Acc_inv x a) ([y:A][h:(R y x)](Acc_rect y (Acc_inv x a y h)))).
End AccRecType.
Definition Acc_rec [P:A->Set] := (Acc_rect P).
(** A relation is well-founded if every element is accessible *)
Definition well_founded := (a:A)(Acc a).
(** well-founded induction on Set and Prop *)
Hypothesis Rwf : well_founded.
Theorem well_founded_induction_type :
(P:A->Type)((x:A)((y:A)(R y x)->(P y))->(P x))->(a:A)(P a).
Proof.
Intros; Apply (Acc_rect P); Auto.
Qed.
Theorem well_founded_induction :
(P:A->Set)((x:A)((y:A)(R y x)->(P y))->(P x))->(a:A)(P a).
Proof.
Exact [P:A->Set](well_founded_induction_type P).
Qed.
Theorem well_founded_ind :
(P:A->Prop)((x:A)((y:A)(R y x)->(P y))->(P x))->(a:A)(P a).
Proof.
Exact [P:A->Prop](well_founded_induction_type P).
Qed.
(** Building fixpoints *)
Section FixPoint.
Variable P : A -> Set.
Variable F : (x:A)((y:A)(R y x)->(P y))->(P x).
Fixpoint Fix_F [x:A;r:(Acc x)] : (P x) :=
(F x [y:A][p:(R y x)](Fix_F y (Acc_inv x r y p))).
Definition fix := [x:A](Fix_F x (Rwf x)).
(** Proof that [well_founded_induction] satisfies the fixpoint equation.
It requires an extra property of the functional *)
Hypothesis F_ext :
(x:A)(f,g:(y:A)(R y x)->(P y))
((y:A)(p:(R y x))((f y p)=(g y p)))->(F x f)=(F x g).
Scheme Acc_inv_dep := Induction for Acc Sort Prop.
Lemma Fix_F_eq
: (x:A)(r:(Acc x))
(F x [y:A][p:(R y x)](Fix_F y (Acc_inv x r y p)))=(Fix_F x r).
Intros x r; Elim r using Acc_inv_dep; Auto.
Qed.
Lemma Fix_F_inv : (x:A)(r,s:(Acc x))(Fix_F x r)=(Fix_F x s).
Intro x; Elim (Rwf x); Intros.
Case (Fix_F_eq x0 r); Case (Fix_F_eq x0 s); Intros.
Apply F_ext; Auto.
Qed.
Lemma Fix_eq : (x:A)(fix x)=(F x [y:A][p:(R y x)](fix y)).
Intro; Unfold fix.
Case (Fix_F_eq x).
Apply F_ext; Intros.
Apply Fix_F_inv.
Qed.
End FixPoint.
End Well_founded.
|