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_tl_eq_eq.ma".
16 include "ground/relocation/pr_isd.ma".
18 (* DIVERGENCE CONDITION FOR PARTIAL RELOCATION MAPS *************************)
20 (* Constructions with pr_eq *************************************************)
22 (*** isdiv_eq_repl_back *)
23 corec lemma pr_isd_eq_repl_back:
24 pr_eq_repl_back … pr_isd.
25 #f1 #H cases (pr_isd_inv_gen … H) -H
26 #g1 #Hg1 #H1 #f2 #Hf cases (pr_eq_inv_next_sn … Hf … H1) -f1
27 /3 width=3 by pr_isd_next/
30 (*** isdiv_eq_repl_fwd *)
31 lemma pr_isd_eq_repl_fwd:
32 pr_eq_repl_fwd … pr_isd.
33 /3 width=3 by pr_isd_eq_repl_back, pr_eq_repl_sym/ qed-.
35 (* Main inversions with pr_eq ***********************************************)
37 (*** isdiv_inv_eq_repl *)
38 corec theorem pr_isd_inv_eq_repl (g1) (g2): 𝛀❨g1❩ → 𝛀❨g2❩ → g1 ≡ g2.
40 cases (pr_isd_inv_gen … H1) -H1
41 cases (pr_isd_inv_gen … H2) -H2
42 /3 width=5 by pr_eq_next/
45 (* Alternative definition with pr_eq ****************************************)
48 corec lemma pr_eq_next_isd (f): ↑f ≡ f → 𝛀❨f❩.
49 #H cases (pr_eq_inv_next_sn … H) -H
50 /4 width=3 by pr_isd_next, pr_eq_trans/
53 (*** eq_next_inv_isdiv *)
54 corec lemma pr_eq_next_inv_isd (g): 𝛀❨g❩ → ↑g ≡ g.
56 /3 width=5 by pr_eq_next/