aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Wellfounded/Lexicographic_Product.v
blob: 01b442a8532cb472ff9f2ff2050fb3793676d802 (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
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
(***********************************************************************)
(*  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*)

(****************************************************************************)
(*                         Bruno Barras  Cristina Cornes                    *)
(*                                                                          *)
(****************************************************************************)

Require Eqdep.
Require Relation_Operators.
Require Transitive_Closure.

(*  From : Constructing Recursion Operators in Type Theory            
           L. Paulson  JSC (1986) 2, 325-355                                
*)
Section WfLexicographic_Product.
Variable A:Set.
Variable B:A->Set.
Variable leA: A->A->Prop.
Variable leB: (x:A)(B x)->(B x)->Prop.


Syntactic Definition LexProd := (lexprod A B leA leB).

Hints Resolve t_step Acc_clos_trans wf_clos_trans.

Lemma acc_A_B_lexprod : (x:A)(Acc A leA x)
        ->((x0:A)(clos_trans A leA x0 x)->(well_founded (B x0) (leB x0)))
                ->(y:(B x))(Acc (B x) (leB x) y)
                        ->(Acc (sigS A B) LexProd (existS A B x y)).
Proof.
 Induction 1; Intros x0 H0 H1 H2 y.
 Induction 1;Intros.
 Apply Acc_intro.
 Induction y0.
 Intros x2 y1 H6.
 Simple Inversion H6;Intros.
 Cut (leA x2 x0);Intros.
 Apply H1;Auto with sets.
 Intros.
 Apply H2.
 Apply t_trans with x2 ;Auto with sets.

 Red in H2.
 Apply H2.
 Auto with sets.

 Injection H8.
 Induction 2.
 Injection H9.
 Induction 2;Auto with sets.

 Elim H8.
 Generalize y2 y' H9 H7 .
 Replace x3 with x0.
 Clear  H7 H8 H9 y2 y' x3 H6 y1 x2 y0.
 Intros.
 Apply H5.
 Elim inj_pair2 with A B x0 y' x1 ;Auto with sets.

 Injection H9;Auto with sets.
Qed.

Theorem wf_lexprod: 
   (well_founded A leA) ->((x:A) (well_founded (B x) (leB x))) 
             -> (well_founded (sigS A B) LexProd).
Proof.
 Intros wfA wfB;Unfold well_founded .
 Induction a;Intros.
 Apply acc_A_B_lexprod;Auto with sets;Intros.
 Red in wfB.
 Auto with sets.
Save.


End WfLexicographic_Product.


Section Wf_Symmetric_Product.
  Variable A:Set.
  Variable B:Set.
  Variable leA: A->A->Prop.
  Variable leB: B->B->Prop.

  Syntactic Definition Symprod := (symprod A B leA leB).

(*i
  Local sig_prod:=
         [x:A*B]<{_:A&B}>Case x of [a:A][b:B](existS A [_:A]B a b) end.

Lemma incl_sym_lexprod: (included (A*B) Symprod  
            (R_o_f (A*B) {_:A&B} sig_prod (lexprod A [_:A]B leA [_:A]leB))).
Proof.
 Red.
 Induction x.
 (Induction y1;Intros).
 Red.
 Unfold sig_prod .
 Inversion_clear H.
 (Apply left_lex;Auto with sets).

 (Apply right_lex;Auto with sets).
Save.
i*)

  Lemma Acc_symprod: (x:A)(Acc A leA x)->(y:B)(Acc B leB y)
                        ->(Acc (A*B) Symprod (x,y)).
Proof.
 Induction 1;Intros.
 Elim H2;Intros.
 Apply Acc_intro;Intros.
 Inversion_clear H5;Auto with sets.
 Apply H1;Auto with sets.
 Apply Acc_intro;Auto with sets.
Save.


Lemma wf_symprod: (well_founded A leA)->(well_founded B leB)
                        ->(well_founded (A*B) Symprod).
Proof.
 Red.
 Induction a;Intros.
 Apply Acc_symprod;Auto with sets.
Save.

End Wf_Symmetric_Product.


Section Swap.

  Variable A:Set.
  Variable R:A->A->Prop.

  Syntactic Definition SwapProd :=(swapprod A R).


  Lemma swap_Acc: (x,y:A)(Acc A*A SwapProd (x,y))->(Acc A*A SwapProd (y,x)).
Proof.
 Intros.
 Inversion_clear H.
 Apply Acc_intro.
 NewDestruct y0;Intros.
 Inversion_clear H;Inversion_clear H1;Apply H0.
 Apply sp_swap.
 Apply right_sym;Auto with sets.

 Apply sp_swap.
 Apply left_sym;Auto with sets.

 Apply sp_noswap.
 Apply right_sym;Auto with sets.

 Apply sp_noswap.
 Apply left_sym;Auto with sets.
Save.


  Lemma Acc_swapprod: (x,y:A)(Acc A R x)->(Acc A R y)
                                ->(Acc A*A SwapProd (x,y)).
Proof.
 Induction 1;Intros.
 Cut (y0:A)(R y0 x0)->(Acc ? SwapProd (y0,y)).
 Clear  H1.
 Elim H2;Intros.
 Cut (y:A)(R y x1)->(Acc ? SwapProd (x0,y)).
 Clear  H3.
 Intro.
 Apply Acc_intro.
 Induction y0;Intros.
 Inversion_clear H5.
 Inversion_clear H6;Auto with sets.

 Apply swap_Acc.
 Inversion_clear H6;Auto with sets.

 Intros.
 Apply H3;Auto with sets;Intros.
 Apply Acc_inv with (y1,x1) ;Auto with sets.
 Apply sp_noswap.
 Apply right_sym;Auto with sets.

 Auto with sets.
Save.


  Lemma wf_swapprod: (well_founded A R)->(well_founded A*A SwapProd).
Proof.
 Red.
 Induction a;Intros.
 Apply Acc_swapprod;Auto with sets.
Save.

End Swap.