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/relocation/pr_coafter_eq.ma".
17 (* RELATIONAL CO-COMPOSITION FOR PARTIAL RELOCATION MAPS ********************)
19 (* Main inversions **********************************************************)
22 corec theorem pr_coafter_mono:
23 ∀f1,f2,x,y. f1 ~⊚ f2 ≘ x → f1 ~⊚ f2 ≘ y → x ≡ y.
24 #f1 #f2 #x #y * -f1 -f2 -x
25 #f1 #f2 #x #g1 [1,2: #g2 ] #g #Hx #H1 [1,2: #H2 ] #H0x #Hy
26 [ cases (pr_coafter_inv_push_bi … Hy … H1 H2) -g1 -g2 /3 width=8 by pr_eq_push/
27 | cases (pr_coafter_inv_push_next … Hy … H1 H2) -g1 -g2 /3 width=8 by pr_eq_next/
28 | cases (pr_coafter_inv_next_sn … Hy … H1) -g1 /3 width=8 by pr_eq_push/
32 (*** coafter_mono_eq *)
33 lemma pr_coafter_mono_eq:
34 ∀f1,f2,f. f1 ~⊚ f2 ≘ f → ∀g1,g2,g. g1 ~⊚ g2 ≘ g →
35 f1 ≡ g1 → f2 ≡ g2 → f ≡ g.
36 /4 width=4 by pr_coafter_mono, pr_coafter_eq_repl_back_dx, pr_coafter_eq_repl_back_sn/ qed-.