1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "Basic_2/unfold/gr2.ma".
17 (* GENERIC RELOCATION WITH PAIRS ********************************************)
19 (* Main properties **********************************************************)
21 theorem at_mono: ∀des,i,i1. @[i] des ≡ i1 → ∀i2. @[i] des ≡ i2 → i1 = i2.
22 #des #i #i1 #H elim H -des -i -i1
23 [ #i #x #H <(at_inv_nil … H) -x //
24 | #des #d #e #i #i1 #Hid #_ #IHi1 #x #H
25 lapply (at_inv_cons_lt … H Hid) -H -Hid /2 width=1/
26 | #des #d #e #i #i1 #Hdi #_ #IHi1 #x #H
27 lapply (at_inv_cons_ge … H Hdi) -H -Hdi /2 width=1/