]> matita.cs.unibo.it Git - helm.git/blob - matita/contribs/POPLmark/Fsub/util.ma
tagged 0.5.0-rc1
[helm.git] / matita / contribs / POPLmark / Fsub / util.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 set "baseuri" "cic:/matita/Fsub/util".
16 include "logic/equality.ma".
17 include "nat/compare.ma".
18 include "list/list.ma".
19 include "list/in.ma".
20
21 (*** useful definitions and lemmas not really related to Fsub ***)
22
23 definition max \def 
24 \lambda m,n.
25   match (leb m n) with
26      [true \Rightarrow n
27      |false \Rightarrow m]. 
28      
29 lemma le_n_max_m_n: \forall m,n:nat. n \le max m n.
30 intros.
31 unfold max.
32 apply (leb_elim m n)
33   [simplify.intros.apply le_n
34   |simplify.intros.autobatch
35   ]
36 qed.
37   
38 lemma le_m_max_m_n: \forall m,n:nat. m \le max m n.
39 intros.
40 unfold max.
41 apply (leb_elim m n)
42   [simplify.intro.assumption
43   |simplify.intros.autobatch
44   ]
45 qed.  
46
47 notation > "hvbox(x ∈ l)"
48   non associative with precedence 30 for @{ 'inlist $x $l }.
49 notation < "hvbox(x \nbsp ∈ \nbsp l)"
50   non associative with precedence 30 for @{ 'inlist $x $l }.
51 interpretation "item in list" 'inlist x l =
52   (cic:/matita/list/in/in_list.ind#xpointer(1/1) _ x l).
53
54 lemma max_case : \forall m,n.(max m n) = match (leb m n) with
55       [ true \Rightarrow n
56       | false \Rightarrow m ].
57 intros;elim m;simplify;reflexivity;
58 qed.