]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama_didactic/bottom.ma
Preparing for 0.5.9 release.
[helm.git] / helm / software / matita / contribs / dama / dama_didactic / bottom.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
16
17 include "nat/times.ma".
18 include "nat/orders.ma".
19
20 include "decl.ma".
21
22 inductive L (T:Type):Type:=
23   bottom: L T
24  |j: T → L T.
25  
26 inductive eq (T:Type) : L T → L T → Prop :=
27  eq_refl:∀x:T. eq T (j ? x) (j ? x).
28
29 notation "hvbox(a break ≡ b)"
30   non associative with precedence 45
31 for @{ 'equiv $a $b }.
32
33 interpretation "uguaglianza parziale" 'equiv x y = (eq ? x y).
34
35 coercion cic:/matita/tests/decl/L.ind#xpointer(1/1/2).
36  
37 lemma sim: ∀T:Type. ∀x,y:T. (j ? x) ≡ (j ? y) → (j ? y) ≡ (j ? x).
38  intros.
39  inversion H.
40  intros.
41  apply eq_refl.
42 qed.
43
44 lemma trans: ∀T:Type. ∀x,y,z:T.
45  (j ? x) ≡ (j ? y) → (j ? y) ≡ (j ? z) → (j ? x) ≡ (j ? z).
46  intros.
47  inversion H1.
48  intros.
49  rewrite > H2 in H.
50  assumption.
51 qed.
52
53 axiom R:Type.
54 axiom R0:R.
55 axiom R1:R.
56 axiom Rplus: L R→L R→L R.
57 axiom Rmult: L R→L R→L R.(*
58 axiom Rdiv: L R→L R→L R.*)
59 axiom Rinv: L R→L R.
60 axiom Relev: L R→L R→L R.
61 axiom Rle: L R→L R→Prop.
62 axiom Rge: L R→L R→Prop.
63
64 interpretation "real plus" 'plus x y = (Rplus x y).
65
66 interpretation "real leq" 'leq x y = (Rle x y).
67
68 interpretation "real geq" 'geq x y = (Rge x y).
69
70 let rec elev (x:L R) (n:nat) on n: L R ≝
71  match n with
72   [O ⇒ match x with [bottom ⇒ bottom ? | j y ⇒ (j ? R1)]
73   | S n ⇒ Rmult x (elev x n)
74   ].
75
76 let rec real_of_nat (n:nat) : L R ≝
77  match n with
78   [ O ⇒ (j ? R0)
79   | S n ⇒ real_of_nat n + (j ? R1)
80   ].
81
82 coercion cic:/matita/tests/decl/real_of_nat.con.
83
84 axiom Rplus_commutative: ∀x,y:R. (j ? x) + (j ? y) ≡ (j ? y) + (j ? x).
85 axiom R0_neutral: ∀x:R. (j ? x) + (j ? R0) ≡ (j ? x).
86 axiom Rmult_commutative: ∀x,y:R. Rmult (j ? x) (j ? y) ≡ Rmult (j ? y) (j ? x).
87 axiom R1_neutral: ∀x:R. Rmult (j ? x) (j ? R1) ≡ (j ? x).
88
89 axiom Rinv_ok:
90  ∀x:R. ¬((j ? x) ≡ (j ? R0)) → Rmult (Rinv (j ? x)) (j ? x) ≡ (j ? R1). 
91 definition is_defined :=
92  λ T:Type. λ x:L T. ∃y:T. x = (j ? y).
93 axiom Rinv_ok2: ∀x:L R. ¬(x = bottom ?) → ¬(x ≡ (j ? R0)) → is_defined ? (Rinv x).
94
95 definition Rdiv :=
96  λ x,y:L R. Rmult x (Rinv y).
97
98 (*
99 lemma pippo: ∀x:R. ¬((j ? x) ≡ (j ? R0)) → Rdiv (j ? R1) (j ? x) ≡ Rinv (j ? x).
100  intros.
101  unfold Rdiv.
102  elim (Rinv_ok2 ? ? H).
103  rewrite > H1.
104  rewrite > Rmult_commutative.
105  apply R1_neutral.
106 *)
107
108 axiom Rdiv_le: ∀x,y:R. (j ? R1) ≤ (j ? y) → Rdiv (j ? x) (j ? y) ≤ (j ? x).
109 axiom R2_1: (j ? R1) ≤ S (S O).
110
111
112 axiom Rdiv_pos: ∀ x,y:R.
113  (j ? R0) ≤ (j ? x) → (j ? R1) ≤ (j ? y) → (j ? R0) ≤ Rdiv (j ? x) (j ? y).
114 axiom Rle_R0_R1: (j ? R0) ≤ (j ? R1).
115 axiom div: ∀x:R. (j ? x) = Rdiv (j ? x) (S (S O)) → (j ? x) = O.