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 (**************************************************************************)
17 include "Track/order.ma".
18 include "NTrack/props.ma".
20 (* Order properties *********************************************************)
22 theorem ntrack_refl: \forall P. \forall c:Formula. \exists r.
23 NTrack P r (pair c c).
24 intros; elim c; clear c;
26 | lapply (insert_total (pair f f1) zero P); [2:autobatch];
27 decompose; autobatch depth = 5 width = 4 size = 8
31 theorem ntrack_trans: \forall p,q,A,B. \forall c:Formula.
32 NTrack leaf p (pair A c) \to NTrack leaf q (pair c B) \to
33 \exists r. NTrack leaf r (pair A B).
35 lapply linear ntrack_track to H as H0;
36 lapply linear ntrack_track to H1 as H;
37 lapply linear track_trans to H0, H as H1; decompose;