(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/Q/q".
+set "baseuri" "cic:/matita/library_auto/Q/q".
include "Z/compare.ma".
include "Z/plus.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/Z/compare".
+set "baseuri" "cic:/matita/library_auto/Z/compare".
include "Z/orders.ma".
include "nat/compare.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/Z/orders".
+set "baseuri" "cic:/matita/library_auto/Z/orders".
include "Z/z.ma".
include "nat/orders.ma".
| (neg m) \Rightarrow m \leq n ]].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "integer 'less or equal to'" 'leq x y = (cic:/matita/Z/orders/Zle.con x y).
+interpretation "integer 'less or equal to'" 'leq x y = (cic:/matita/library_auto/Z/orders/Zle.con x y).
(*CSC: the URI must disappear: there is a bug now *)
interpretation "integer 'neither less nor equal to'" 'nleq x y =
- (cic:/matita/logic/connectives/Not.con (cic:/matita/Z/orders/Zle.con x y)).
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/library_auto/Z/orders/Zle.con x y)).
definition Zlt : Z \to Z \to Prop \def
\lambda x,y:Z.
| (neg m) \Rightarrow m<n ]].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "integer 'less than'" 'lt x y = (cic:/matita/Z/orders/Zlt.con x y).
+interpretation "integer 'less than'" 'lt x y = (cic:/matita/library_auto/Z/orders/Zlt.con x y).
(*CSC: the URI must disappear: there is a bug now *)
interpretation "integer 'not less than'" 'nless x y =
- (cic:/matita/logic/connectives/Not.con (cic:/matita/Z/orders/Zlt.con x y)).
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/library_auto/Z/orders/Zlt.con x y)).
theorem irreflexive_Zlt: irreflexive Z Zlt.
unfold irreflexive.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/Z/plus".
+set "baseuri" "cic:/matita/library_auto/Z/plus".
include "Z/z.ma".
include "nat/minus.ma".
| (neg n) \Rightarrow (neg (pred ((S m)+(S n))))] ].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "integer plus" 'plus x y = (cic:/matita/Z/plus/Zplus.con x y).
+interpretation "integer plus" 'plus x y = (cic:/matita/library_auto/Z/plus/Zplus.con x y).
theorem Zplus_z_OZ: \forall z:Z. z+OZ = z.
intro.
| (neg n) \Rightarrow (pos n) ].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "integer unary minus" 'uminus x = (cic:/matita/Z/plus/Zopp.con x).
+interpretation "integer unary minus" 'uminus x = (cic:/matita/library_auto/Z/plus/Zopp.con x).
theorem Zopp_Zplus: \forall x,y:Z. -(x+y) = -x + -y.
intros.
(* minus *)
definition Zminus : Z \to Z \to Z \def \lambda x,y:Z. x + (-y).
-interpretation "integer minus" 'minus x y = (cic:/matita/Z/plus/Zminus.con x y).
+interpretation "integer minus" 'minus x y = (cic:/matita/library_auto/Z/plus/Zminus.con x y).
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/Z/times".
+set "baseuri" "cic:/matita/library_auto/Z/times".
include "nat/lt_arith.ma".
include "Z/plus.ma".
| (neg n) \Rightarrow (pos (pred ((S m) * (S n))))]].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "integer times" 'times x y = (cic:/matita/Z/times/Ztimes.con x y).
+interpretation "integer times" 'times x y = (cic:/matita/library_auto/Z/times/Ztimes.con x y).
theorem Ztimes_z_OZ: \forall z:Z. z*OZ = OZ.
intro.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/Z/z".
+set "baseuri" "cic:/matita/library_auto/Z/z".
include "datatypes/bool.ma".
include "nat/nat.ma".
[ O \Rightarrow OZ
| (S n)\Rightarrow pos n].
-coercion cic:/matita/Z/z/Z_of_nat.con.
+coercion cic:/matita/library_auto/Z/z/Z_of_nat.con.
definition neg_Z_of_nat \def
\lambda n. match n with
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/chinese_reminder".
+set "baseuri" "cic:/matita/library_auto/nat/chinese_reminder".
include "nat/exp.ma".
include "nat/gcd.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/compare".
+set "baseuri" "cic:/matita/library_auto/nat/compare".
include "datatypes/bool.ma".
include "datatypes/compare.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/congruence".
+set "baseuri" "cic:/matita/library_auto/nat/congruence".
include "nat/relevant_equations.ma".
include "nat/primes.ma".
\lambda n,m,p:nat. mod n p = mod m p.
interpretation "congruent" 'congruent n m p =
- (cic:/matita/nat/congruence/congruent.con n m p).
+ (cic:/matita/library_auto/nat/congruence/congruent.con n m p).
notation < "hvbox(n break \cong\sub p m)"
(*non associative*) with precedence 45
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/count".
+set "baseuri" "cic:/matita/library_auto/nat/count".
include "nat/relevant_equations.ma".
include "nat/sigma_and_pi.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/div_and_mod".
+set "baseuri" "cic:/matita/library_auto/nat/div_and_mod".
include "datatypes/constructors.ma".
include "nat/minus.ma".
| (S p) \Rightarrow mod_aux n n p].
interpretation "natural remainder" 'module x y =
- (cic:/matita/nat/div_and_mod/mod.con x y).
+ (cic:/matita/library_auto/nat/div_and_mod/mod.con x y).
let rec div_aux p m n : nat \def
match (leb m n) with
| (S p) \Rightarrow div_aux n n p].
interpretation "natural divide" 'divide x y =
- (cic:/matita/nat/div_and_mod/div.con x y).
+ (cic:/matita/library_auto/nat/div_and_mod/div.con x y).
theorem le_mod_aux_m_m:
\forall p,n,m. n \leq p \to (mod_aux p n m) \leq m.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/euler_theorem".
+set "baseuri" "cic:/matita/library_auto/nat/euler_theorem".
include "nat/map_iter_p.ma".
include "nat/totient.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/exp".
+set "baseuri" "cic:/matita/library_auto/nat/exp".
include "nat/div_and_mod.ma".
[ O \Rightarrow (S O)
| (S p) \Rightarrow (times n (exp n p)) ].
-interpretation "natural exponent" 'exp a b = (cic:/matita/nat/exp/exp.con a b).
+interpretation "natural exponent" 'exp a b = (cic:/matita/library_auto/nat/exp/exp.con a b).
theorem exp_plus_times : \forall n,p,q:nat.
n \sup (p + q) = (n \sup p) * (n \sup q).
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/factorial".
+set "baseuri" "cic:/matita/library_auto/nat/factorial".
include "nat/le_arith.ma".
[ O \Rightarrow (S O)
| (S m) \Rightarrow (S m)*(fact m)].
-interpretation "factorial" 'fact n = (cic:/matita/nat/factorial/fact.con n).
+interpretation "factorial" 'fact n = (cic:/matita/library_auto/nat/factorial/fact.con n).
theorem le_SO_fact : \forall n. (S O) \le n!.
intro.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/factorization".
+set "baseuri" "cic:/matita/library_auto/nat/factorization".
include "nat/ord.ma".
include "nat/gcd.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/fermat_little_theorem".
+set "baseuri" "cic:/matita/library_auto/nat/fermat_little_theorem".
include "nat/exp.ma".
include "nat/gcd.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/gcd".
+set "baseuri" "cic:/matita/library_auto/nat/gcd".
include "nat/primes.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/le_arith".
+set "baseuri" "cic:/matita/library_auto/nat/le_arith".
include "nat/times.ma".
include "nat/orders.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/lt_arith".
+set "baseuri" "cic:/matita/library_auto/nat/lt_arith".
include "nat/div_and_mod.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/map_iter_p.ma".
+set "baseuri" "cic:/matita/library_auto/nat/map_iter_p.ma".
include "nat/permutation.ma".
include "nat/count.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/minimization".
+set "baseuri" "cic:/matita/library_auto/nat/minimization".
include "nat/minus.ma".
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/minus".
+set "baseuri" "cic:/matita/library_auto/nat/minus".
include "nat/le_arith.ma".
include "nat/compare.ma".
| (S q) \Rightarrow minus p q ]].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural minus" 'minus x y = (cic:/matita/nat/minus/minus.con x y).
+interpretation "natural minus" 'minus x y = (cic:/matita/library_auto/nat/minus/minus.con x y).
theorem minus_n_O: \forall n:nat.n=n-O.
intros.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/nat".
+set "baseuri" "cic:/matita/library_auto/nat/nat".
include "higher_order_defs/functions.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/nth_prime".
+set "baseuri" "cic:/matita/library_auto/nat/nth_prime".
include "nat/primes.ma".
include "nat/lt_arith.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/ord".
+set "baseuri" "cic:/matita/library_auto/nat/ord".
include "datatypes/constructors.ma".
include "nat/exp.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/orders".
+set "baseuri" "cic:/matita/library_auto/nat/orders".
include "nat/nat.ma".
include "higher_order_defs/ordering.ma".
| le_S : \forall m:nat. le n m \to le n (S m).
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural 'less or equal to'" 'leq x y = (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y).
+interpretation "natural 'less or equal to'" 'leq x y = (cic:/matita/library_auto/nat/orders/le.ind#xpointer(1/1) x y).
(*CSC: the URI must disappear: there is a bug now *)
interpretation "natural 'neither less nor equal to'" 'nleq x y =
(cic:/matita/logic/connectives/Not.con
- (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y)).
+ (cic:/matita/library_auto/nat/orders/le.ind#xpointer(1/1) x y)).
definition lt: nat \to nat \to Prop \def
\lambda n,m:nat.(S n) \leq m.
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural 'less than'" 'lt x y = (cic:/matita/nat/orders/lt.con x y).
+interpretation "natural 'less than'" 'lt x y = (cic:/matita/library_auto/nat/orders/lt.con x y).
(*CSC: the URI must disappear: there is a bug now *)
interpretation "natural 'not less than'" 'nless x y =
- (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/lt.con x y)).
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/library_auto/nat/orders/lt.con x y)).
definition ge: nat \to nat \to Prop \def
\lambda n,m:nat.m \leq n.
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural 'greater or equal to'" 'geq x y = (cic:/matita/nat/orders/ge.con x y).
+interpretation "natural 'greater or equal to'" 'geq x y = (cic:/matita/library_auto/nat/orders/ge.con x y).
definition gt: nat \to nat \to Prop \def
\lambda n,m:nat.m<n.
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural 'greater than'" 'gt x y = (cic:/matita/nat/orders/gt.con x y).
+interpretation "natural 'greater than'" 'gt x y = (cic:/matita/library_auto/nat/orders/gt.con x y).
(*CSC: the URI must disappear: there is a bug now *)
interpretation "natural 'not greater than'" 'ngtr x y =
- (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/gt.con x y)).
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/library_auto/nat/orders/gt.con x y)).
theorem transitive_le : transitive nat le.
unfold transitive.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/permutation".
+set "baseuri" "cic:/matita/library_auto/nat/permutation".
include "nat/compare.ma".
include "nat/sigma_and_pi.ma".
for @{ 'transposition $i $j $n}.
interpretation "natural transposition" 'transposition i j n =
- (cic:/matita/nat/permutation/transpose.con i j n).
+ (cic:/matita/library_auto/nat/permutation/transpose.con i j n).
lemma transpose_i_j_i: \forall i,j:nat. transpose i j i = j.
intros.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/plus".
+set "baseuri" "cic:/matita/library_auto/nat/plus".
include "nat/nat.ma".
| (S p) \Rightarrow S (plus p m) ].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural plus" 'plus x y = (cic:/matita/nat/plus/plus.con x y).
+interpretation "natural plus" 'plus x y = (cic:/matita/library_auto/nat/plus/plus.con x y).
theorem plus_n_O: \forall n:nat. n = n+O.
intros.elim n
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/primes".
+set "baseuri" "cic:/matita/library_auto/nat/primes".
include "nat/div_and_mod.ma".
include "nat/minimization.ma".
inductive divides (n,m:nat) : Prop \def
witness : \forall p:nat.m = times n p \to divides n m.
-interpretation "divides" 'divides n m = (cic:/matita/nat/primes/divides.ind#xpointer(1/1) n m).
+interpretation "divides" 'divides n m = (cic:/matita/library_auto/nat/primes/divides.ind#xpointer(1/1) n m).
interpretation "not divides" 'ndivides n m =
- (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/primes/divides.ind#xpointer(1/1) n m)).
+ (cic:/matita/logic/connectives/Not.con (cic:/matita/library_auto/nat/primes/divides.ind#xpointer(1/1) n m)).
theorem reflexive_divides : reflexive nat divides.
unfold reflexive.
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/relevant_equations".
+set "baseuri" "cic:/matita/library_auto/nat/relevant_equations".
include "nat/times.ma".
include "nat/minus.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/sigma_and_pi".
+set "baseuri" "cic:/matita/library_auto/nat/sigma_and_pi".
include "nat/factorial.ma".
include "nat/exp.ma".
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/times".
+set "baseuri" "cic:/matita/library_auto/nat/times".
include "nat/plus.ma".
| (S p) \Rightarrow m+(times p m) ].
(*CSC: the URI must disappear: there is a bug now *)
-interpretation "natural times" 'times x y = (cic:/matita/nat/times/times.con x y).
+interpretation "natural times" 'times x y = (cic:/matita/library_auto/nat/times/times.con x y).
theorem times_n_O: \forall n:nat. O = n*O.
intros.elim n
(* *)
(**************************************************************************)
-set "baseuri" "cic:/matita/nat/totient".
+set "baseuri" "cic:/matita/library_auto/nat/totient".
include "nat/count.ma".
include "nat/chinese_reminder.ma".