]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user andrea
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Sat, 27 Apr 2013 14:59:25 +0000 (14:59 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Sat, 27 Apr 2013 14:59:25 +0000 (14:59 +0000)
weblib/arithmetics/bigO.ma [new file with mode: 0644]
weblib/basics/sets.ma [new file with mode: 0644]

diff --git a/weblib/arithmetics/bigO.ma b/weblib/arithmetics/bigO.ma
new file mode 100644 (file)
index 0000000..95ccd1f
--- /dev/null
@@ -0,0 +1,60 @@
+include "arithmetics/nat.ma".
+include "basics/sets.ma".
+
+(*  O f g means g ∈ O(f) *)
+definition O: relation (nat→nat) ≝
+  λf,g. ∃c.∃n0.∀n. n0 ≤ n → g n ≤ c* (f n).
+  
+lemma O_refl: ∀s. O s s.
+#s %{1} %{0} #n #_ >commutative_times <times_n_1 @le_n qed.
+
+lemma O_trans: ∀s1,s2,s3. O s2 s1 → O s3 s2 → O s3 s1. 
+#s1 #s2 #s3 * #c1 * #n1 #H1 * #c2 * # n2 #H2 %{(c1*c2)}
+%{(max n1 n2)} #n #Hmax 
+@(transitive_le … (H1 ??)) [@(le_maxl … Hmax)]
+>associative_times @le_times [//|@H2 @(le_maxr … Hmax)]
+qed.
+
+lemma sub_O_to_O: ∀s1,s2. O s1 ⊆ O s2 → O s2 s1.
+#s1 #s2 #H @H // qed.
+
+lemma O_to_sub_O: ∀s1,s2. O s2 s1 → O s1 ⊆ O s2.
+#s1 #s2 #H #g #Hg @(O_trans … H) // qed. 
+
+definition sum_f ≝ λf,g:nat→nat.λn.f n + g n.
+interpretation "function sum" 'plus f g = (sum_f f g).
+
+lemma O_plus: ∀f,g,s. O s f → O s g → O s (f+g).
+#f #g #s * #cf * #nf #Hf * #cg * #ng #Hg
+%{(cf+cg)} %{(max nf ng)} #n #Hmax normalize 
+>distributive_times_plus_r @le_plus 
+  [@Hf @(le_maxl … Hmax) |@Hg @(le_maxr … Hmax) ]
+qed.
+lemma O_plus_l: ∀f,s1,s2. O s1 f → O (s1+s2) f.
+#f #s1 #s2 * #c * #a #Os1f %{c} %{a} #n #lean 
+@(transitive_le … (Os1f n lean)) @le_times //
+qed.
+
+lemma O_absorbl: ∀f,g,s. O s f → O f g → O s (g+f).
+#f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) //
+qed.
+
+lemma O_absorbr: ∀f,g,s. O s f → O f g → O s (f+g).
+#f #g #s #Osf #Ofg @(O_plus … Osf) @(O_trans … Osf) //
+qed.
+
+(* 
+lemma O_ff: ∀f,s. O s f → O s (f+f).
+#f #s #Osf /2/ 
+qed. *)
+
+lemma O_ext2: ∀f,g,s. O s f → (∀x.f x = g x) → O s g.
+#f #g #s * #c * #a #Osf #eqfg %{c} %{a} #n #lean <eqfg @Osf //
+qed.    
+
+
+definition not_O ≝ λf,g.∀c,n0.∃n. n0 ≤ n ∧ c* (f n) < g n .
+
+(* this is the only classical result *)
+axiom not_O_def: ∀f,g. ¬ O f g → not_O f g.
\ No newline at end of file
diff --git a/weblib/basics/sets.ma b/weblib/basics/sets.ma
new file mode 100644 (file)
index 0000000..d441706
--- /dev/null
@@ -0,0 +1,108 @@
+include "basics/logic.ma".
+
+(**** a subset of A is just an object of type A→Prop ****)
+
+\ 5img class="anchor" src="icons/tick.png" id="empty_set"\ 6definition empty_set ≝ λA:Type[0].λa:A.\ 5a href="cic:/matita/basics/logic/False.ind(1,0,0)"\ 6False\ 5/a\ 6.
+notation "\emptyv" non associative with precedence 90 for @{'empty_set}.
+interpretation "empty set" 'empty_set = (empty_set ?).
+
+\ 5img class="anchor" src="icons/tick.png" id="singleton"\ 6definition singleton ≝ λA.λx,a:A.x\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6a.
+(* notation "{x}" non associative with precedence 90 for @{'sing_lang $x}. *)
+interpretation "singleton" 'singl x = (singleton ? x).
+
+\ 5img class="anchor" src="icons/tick.png" id="union"\ 6definition union : ∀A:Type[0].∀P,Q.A → Prop ≝ λA,P,Q,a.P a \ 5a title="logical or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 Q a.
+interpretation "union" 'union a b = (union ? a b).
+
+\ 5img class="anchor" src="icons/tick.png" id="intersection"\ 6definition intersection : ∀A:Type[0].∀P,Q.A→Prop ≝ λA,P,Q,a.P a \ 5a title="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 Q a.
+interpretation "intersection" 'intersects a b = (intersection ? a b).
+
+\ 5img class="anchor" src="icons/tick.png" id="complement"\ 6definition complement ≝ λU:Type[0].λA:U → Prop.λw.\ 5a title="logical not" href="cic:/fakeuri.def(1)"\ 6¬\ 5/a\ 6 A w.
+interpretation "complement" 'not a = (complement ? a).
+
+\ 5img class="anchor" src="icons/tick.png" id="substraction"\ 6definition substraction := λU:Type[0].λA,B:U → Prop.λw.A w \ 5a title="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="logical not" href="cic:/fakeuri.def(1)"\ 6¬\ 5/a\ 6 B w.
+interpretation "substraction" 'minus a b = (substraction ? a b).
+
+\ 5img class="anchor" src="icons/tick.png" id="subset"\ 6definition subset: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ λA,P,Q.∀a:A.(P a → Q a).
+interpretation "subset" 'subseteq a b = (subset ? a b).
+
+(* extensional equality *)
+\ 5img class="anchor" src="icons/tick.png" id="eqP"\ 6definition eqP ≝ λA:Type[0].λP,Q:A → Prop.∀a:A.P a \ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 Q a.
+(* ≐ is typed as \doteq *)
+notation "A ≐ B" non associative with precedence 45 for @{'eqP $A $B}.
+interpretation "extensional equality" 'eqP a b = (eqP ? a b).
+
+\ 5img class="anchor" src="icons/tick.png" id="eqP_sym"\ 6lemma eqP_sym: ∀U.∀A,B:U →Prop. 
+  A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B → B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A.
+#U #A #B #eqAB #a @\ 5a href="cic:/matita/basics/logic/iff_sym.def(2)"\ 6iff_sym\ 5/a\ 6 @eqAB qed.
+\ 5img class="anchor" src="icons/tick.png" id="eqP_trans"\ 6lemma eqP_trans: ∀U.∀A,B,C:U →Prop. 
+  A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B → B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C → A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C.
+#U #A #B #C #eqAB #eqBC #a @\ 5a href="cic:/matita/basics/logic/iff_trans.def(2)"\ 6iff_trans\ 5/a\ 6 // qed.
+
+\ 5img class="anchor" src="icons/tick.png" id="eqP_union_r"\ 6lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. 
+  A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C  → A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B.
+#U #A #B #C #eqAB #a @\ 5a href="cic:/matita/basics/logic/iff_or_r.def(2)"\ 6iff_or_r\ 5/a\ 6 @eqAB qed.
+  
+\ 5img class="anchor" src="icons/tick.png" id="eqP_union_l"\ 6lemma eqP_union_l: ∀U.∀A,B,C:U →Prop. 
+  B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C  → A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C.
+#U #A #B #C #eqBC #a @\ 5a href="cic:/matita/basics/logic/iff_or_l.def(2)"\ 6iff_or_l\ 5/a\ 6 @eqBC qed.
+  
+\ 5img class="anchor" src="icons/tick.png" id="eqP_intersect_r"\ 6lemma eqP_intersect_r: ∀U.∀A,B,C:U →Prop. 
+  A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C  → A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B.
+#U #A #B #C #eqAB #a @\ 5a href="cic:/matita/basics/logic/iff_and_r.def(2)"\ 6iff_and_r\ 5/a\ 6 @eqAB qed.
+  
+\ 5img class="anchor" src="icons/tick.png" id="eqP_intersect_l"\ 6lemma eqP_intersect_l: ∀U.∀A,B,C:U →Prop. 
+  B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C  → A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C.
+#U #A #B #C #eqBC #a @\ 5a href="cic:/matita/basics/logic/iff_and_l.def(2)"\ 6iff_and_l\ 5/a\ 6 @eqBC qed.
+
+\ 5img class="anchor" src="icons/tick.png" id="eqP_substract_r"\ 6lemma eqP_substract_r: ∀U.∀A,B,C:U →Prop. 
+  A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C  → A \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 B.
+#U #A #B #C #eqAB #a @\ 5a href="cic:/matita/basics/logic/iff_and_r.def(2)"\ 6iff_and_r\ 5/a\ 6 @eqAB qed.
+  
+\ 5img class="anchor" src="icons/tick.png" id="eqP_substract_l"\ 6lemma eqP_substract_l: ∀U.∀A,B,C:U →Prop. 
+  B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C  → A \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 C.
+#U #A #B #C #eqBC #a @\ 5a href="cic:/matita/basics/logic/iff_and_l.def(2)"\ 6iff_and_l\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/iff_not.def(4)"\ 6iff_not\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
+
+(* set equalities *)
+\ 5img class="anchor" src="icons/tick.png" id="union_empty_r"\ 6lemma union_empty_r: ∀U.∀A:U→Prop. 
+  A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="empty set" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A.
+#U #A #w % [* // normalize #abs @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5/span\ 6\ 5/span\ 6/ | /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/]
+qed.
+
+\ 5img class="anchor" src="icons/tick.png" id="union_comm"\ 6lemma union_comm : ∀U.∀A,B:U →Prop. 
+  A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A.
+#U #A #B #a % * /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed. 
+
+\ 5img class="anchor" src="icons/tick.png" id="union_assoc"\ 6lemma union_assoc: ∀U.∀A,B,C:U → Prop. 
+  A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (B \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C).
+#S #A #B #C #w % [* [* /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ | /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] | * [/\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ | * /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/]
+qed.   
+
+\ 5img class="anchor" src="icons/tick.png" id="cap_comm"\ 6lemma cap_comm : ∀U.∀A,B:U →Prop. 
+  A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A.
+#U #A #B #a % * /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed. 
+
+\ 5img class="anchor" src="icons/tick.png" id="union_idemp"\ 6lemma union_idemp: ∀U.∀A:U →Prop. 
+  A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A.
+#U #A #a % [* // | /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] qed. 
+
+\ 5img class="anchor" src="icons/tick.png" id="cap_idemp"\ 6lemma cap_idemp: ∀U.∀A:U →Prop. 
+  A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A.
+#U #A #a % [* // | /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] qed. 
+
+(*distributivities *)
+
+\ 5img class="anchor" src="icons/tick.png" id="distribute_intersect"\ 6lemma distribute_intersect : ∀U.∀A,B,C:U→Prop. 
+  (A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B) \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C) \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (B \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C).
+#U #A #B #C #w % [* * /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ | * * /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] 
+qed.
+  
+\ 5img class="anchor" src="icons/tick.png" id="distribute_substract"\ 6lemma distribute_substract : ∀U.∀A,B,C:U→Prop. 
+  (A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B) \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 C \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (A \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 C) \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (B \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 C).
+#U #A #B #C #w % [* * /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ | * * /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] 
+qed.
+
+(* substraction *)
+\ 5img class="anchor" src="icons/tick.png" id="substract_def"\ 6lemma substract_def:∀U.∀A,B:U→Prop. A\ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6\ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="complement" href="cic:/fakeuri.def(1)"\ 6¬\ 5/a\ 6B.
+#U #A #B #w normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
+qed.
\ No newline at end of file