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/notation/functions/upspoon_1.ma".
16 include "ground/lib/stream.ma".
17 include "ground/arith/pnat.ma".
19 (* RELOCATION P-STREAM ******************************************************)
21 (* Basic inversion lemmas ***************************************************)
23 lemma eq_inv_gr_push_bi: injective ? ? gr_push.
24 #f1 #f2 <gr_push_unfold <gr_push_unfold #H destruct //
27 lemma eq_inv_gr_push_next: ∀f1,f2. ⫯f1 = ↑f2 → ⊥.
28 #f1 * #p2 #f2 <gr_push_unfold <gr_next_unfold #H destruct
31 lemma eq_inv_gr_next_push: ∀f1,f2. ↑f1 = ⫯f2 → ⊥.
32 * #p1 #f1 #f2 <gr_next_unfold <gr_push_unfold #H destruct
35 lemma eq_inv_gr_next_bi: injective ? ? gr_next.
36 * #p1 #f1 * #p2 #f2 <gr_next_unfold <gr_next_unfold #H destruct //
39 lemma case_prop (Q:predicate gr_map):
40 (∀f. Q (⫯f)) → (∀f. Q (↑f)) → ∀f. Q f.
44 lemma case_type0 (Q:gr_map→Type[0]):
45 (∀f. Q (⫯f)) → (∀f. Q (↑f)) → ∀f. Q f.
49 lemma iota_push: ∀Q,a,b,f. a f = case_type0 Q a b (⫯f).
52 lemma iota_next: ∀Q,a,b,f. b f = case_type0 Q a b (↑f).