]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/tests/ng_lexiconn.ma
Preparing for 0.5.9 release.
[helm.git] / helm / software / matita / tests / ng_lexiconn.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ng_pts.ma".
16
17 axiom A : Type.
18
19 axiom a : A. 
20 axiom b : A.
21
22 axiom B : Type.
23
24 axiom c : B. 
25 axiom d : B.
26
27
28 notation "#" with precedence 90 for @{ 'foo }.
29 interpretation "a" 'foo = a. 
30 interpretation "b" 'foo = b. 
31 interpretation "c" 'foo = c. 
32 interpretation "d" 'foo = d. 
33
34 alias symbol "foo" = "c".
35 alias symbol "foo" = "b".
36 alias symbol "foo" = "d".
37 alias symbol "foo" = "b".   
38
39 lemma xx : ∀P: A → B → Prop. P # #. (* NON STAMPA GLI ALIAS *) 
40