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.ma".
16 include "ground/relocation/pr_sand.ma".
18 (* RELATIONAL INTERSECTION FOR PARTIAL RELOCATION MAPS **********************)
20 (* Constructions with pr_eq *************************************************)
22 (*** sand_eq_repl_back1 *)
23 corec lemma pr_sand_eq_repl_back_sn:
24 ∀f2,f. pr_eq_repl_back … (λf1. f1 ⋒ f2 ≘ f).
25 #f2 #f #f1 * -f1 -f2 -f
26 #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0 #x #Hx
27 try cases (pr_eq_inv_push_sn … Hx … H1) try cases (pr_eq_inv_next_sn … Hx … H1) -g1
28 /3 width=7 by pr_sand_push_bi, pr_sand_next_push, pr_sand_push_next, pr_sand_next_bi/
31 (*** sand_eq_repl_fwd1 *)
32 lemma pr_sand_eq_repl_fwd_sn:
33 ∀f2,f. pr_eq_repl_fwd … (λf1. f1 ⋒ f2 ≘ f).
34 #f2 #f @pr_eq_repl_sym /2 width=3 by pr_sand_eq_repl_back_sn/
37 (*** sand_eq_repl_back2 *)
38 corec lemma pr_sand_eq_repl_back_dx:
39 ∀f1,f. pr_eq_repl_back … (λf2. f1 ⋒ f2 ≘ f).
40 #f1 #f #f2 * -f1 -f2 -f
41 #f1 #f2 #f #g1 #g2 #g #Hf #H #H2 #H0 #x #Hx
42 try cases (pr_eq_inv_push_sn … Hx … H2) try cases (pr_eq_inv_next_sn … Hx … H2) -g2
43 /3 width=7 by pr_sand_push_bi, pr_sand_next_push, pr_sand_push_next, pr_sand_next_bi/
46 (*** sand_eq_repl_fwd2 *)
47 lemma sand_eq_repl_fwd_dx:
48 ∀f1,f. pr_eq_repl_fwd … (λf2. f1 ⋒ f2 ≘ f).
49 #f1 #f @pr_eq_repl_sym /2 width=3 by pr_sand_eq_repl_back_dx/
52 (*** sand_eq_repl_back3 *)
53 corec lemma pr_sand_eq_repl_back:
54 ∀f1,f2. pr_eq_repl_back … (λf. f1 ⋒ f2 ≘ f).
55 #f1 #f2 #f * -f1 -f2 -f
56 #f1 #f2 #f #g1 #g2 #g #Hf #H #H2 #H0 #x #Hx
57 try cases (pr_eq_inv_push_sn … Hx … H0) try cases (pr_eq_inv_next_sn … Hx … H0) -g
58 /3 width=7 by pr_sand_push_bi, pr_sand_next_push, pr_sand_push_next, pr_sand_next_bi/
61 (*** sand_eq_repl_fwd3 *)
62 lemma pr_sand_eq_repl_fwd:
63 ∀f1,f2. pr_eq_repl_fwd … (λf. f1 ⋒ f2 ≘ f).
64 #f1 #f2 @pr_eq_repl_sym /2 width=3 by pr_sand_eq_repl_back/