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_sle_sle.ma".
16 include "ground/relocation/pr_sor.ma".
18 (* RELATIONAL UNION FOR PARTIAL RELOCATION MAPS *****************************)
20 (* Inversions with pr_sle ***************************************************)
22 (*** sor_inv_sle_sn *)
23 corec lemma pr_sor_inv_sle_sn:
24 ∀f1,f2,f. f1 ⋓ f2 ≘ f → f1 ⊆ f.
25 #f1 #f2 #f * -f1 -f2 -f
26 #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0
27 /3 width=5 by pr_sle_push, pr_sle_next, pr_sle_weak/
30 (*** sor_inv_sle_dx *)
31 corec lemma pr_sor_inv_sle_dx:
32 ∀f1,f2,f. f1 ⋓ f2 ≘ f → f2 ⊆ f.
33 #f1 #f2 #f * -f1 -f2 -f
34 #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0
35 /3 width=5 by pr_sle_push, pr_sle_next, pr_sle_weak/
38 (*** sor_inv_sle_sn_trans *)
39 lemma pr_sor_inv_sle_sn_trans:
40 ∀f1,f2,f. f1 ⋓ f2 ≘ f → ∀g. g ⊆ f1 → g ⊆ f.
41 /3 width=4 by pr_sor_inv_sle_sn, pr_sle_trans/ qed-.
43 (*** sor_inv_sle_dx_trans *)
44 lemma pr_sor_inv_sle_dx_trans:
45 ∀f1,f2,f. f1 ⋓ f2 ≘ f → ∀g. g ⊆ f2 → g ⊆ f.
46 /3 width=4 by pr_sor_inv_sle_dx, pr_sle_trans/ qed-.
49 axiom pr_sor_inv_sle_bi:
50 ∀f1,f2,f. f1 ⋓ f2 ≘ f → ∀g. f1 ⊆ g → f2 ⊆ g → f ⊆ g.
52 (* Constructions with pr_sle ************************************************)
55 corec lemma pr_sor_sle_dx:
56 ∀f1,f2. f1 ⊆ f2 → f1 ⋓ f2 ≘ f2.
58 /3 width=7 by pr_sor_push_bi, pr_sor_next_bi, pr_sor_push_next/
62 corec lemma pr_sor_sle_sn:
63 ∀f1,f2. f1 ⊆ f2 → f2 ⋓ f1 ≘ f2.
65 /3 width=7 by pr_sor_push_bi, pr_sor_next_bi, pr_sor_next_push/