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 "ground_2/notation/relations/rplusminus_4.ma".
16 include "ground_2/ynat/ynat_plus.ma".
18 (* NATURAL NUMBERS WITH INFINITY ********************************************)
20 (* algebraic x + y1 - y2 = z *)
21 inductive yrpm (x:ynat) (y1:ynat) (y2:ynat): predicate ynat ≝
22 | yrpm_ge: y2 ≤ y1 → yrpm x y1 y2 (x + (y1 - y2))
23 | yrpm_lt: y1 < y2 → yrpm x y1 y2 (x - (y2 - y1))
26 interpretation "ynat 'algebraic plus-minus' (relational)"
27 'RPlusMinus x y1 y2 z = (yrpm x y1 y2 z).
29 (* Basic inversion lemmas ***************************************************)
31 lemma ypm_inv_ge: ∀x,y1,y2,z. x ⊞ y1 ⊟ y2 ≡ z →
32 y2 ≤ y1 → z = x + (y1 - y2).
34 #Hy12 #H elim (ylt_yle_false … H) -H //
37 lemma ypm_inv_lt: ∀x,y1,y2,z. x ⊞ y1 ⊟ y2 ≡ z →
38 y1 < y2 → z = x - (y2 - y1).
40 #Hy21 #H elim (ylt_yle_false … H) -H //
43 (* Advanced inversion lemmas ************************************************)
45 lemma ypm_inv_le: ∀x,y1,y2,z. x ⊞ y1 ⊟ y2 ≡ z →
46 y1 ≤ y2 → z = x - (y2 - y1).
47 #x #y1 #y2 #z #H #Hy12 elim (yle_split_eq … Hy12) -Hy12 #Hy12
48 [ /2 width=1 by ypm_inv_lt/
49 | >(ypm_inv_ge … H) -H // destruct >yminus_refl //