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_isi_pat.ma".
16 include "ground/relocation/pr_ist.ma".
18 (* TOTALITY CONDITION FOR PARTIAL RELOCATION MAPS ***************************)
20 (* Advanced constructions with pr_isi ***************************************)
23 lemma pr_isi_pat_total: ∀f. 𝐓❨f❩ → (∀i1,i2. @❨i1,f❩ ≘ i2 → i1 = i2) → 𝐈❨f❩.
24 #f #H1f #H2f @pr_isi_pat
25 #i lapply (H1f i) -H1f *
26 #j #Hf >(H2f … Hf) in ⊢ (???%); -H2f //