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