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 "delayed_updating/syntax/prototerm.ma".
16 include "ground/lib/subset_equivalence.ma".
18 (* EQUIVALENCE FOR PROTOTERM ************************************************)
20 (* Constructions with prototerm_grafted *************************************)
22 lemma grafted_empty_dx (t):
27 (* Constructions with prototerm_root ****************************************)
29 lemma prototerm_root_incl_repl:
30 ∀t1,t2. t1 ⊆ t2 → ▵t1 ⊆ ▵t2.
31 #t1 #t2 #Ht #p * #q #Hq
32 /3 width=2 by ex_intro/
35 lemma prototerm_root_eq_repl:
36 ∀t1,t2. t1 ⇔ t2 → ▵t1 ⇔ ▵t2.
38 /3 width=3 by conj, prototerm_root_incl_repl/