]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user mkmluser
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 10 Jul 2012 08:54:01 +0000 (08:54 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 10 Jul 2012 08:54:01 +0000 (08:54 +0000)
weblib/cicm2012/disambiguation.ma [new file with mode: 0644]
weblib/cicm2012/naturals.ma [new file with mode: 0644]
weblib/cicm2012/part1.ma [new file with mode: 0644]
weblib/cicm2012/part2.ma [new file with mode: 0644]
weblib/cicm2012/rationals.ma [new file with mode: 0644]
weblib/cicm2012/reals.ma [new file with mode: 0644]
weblib/commit_test2.ma
weblib/part1.ma [new file with mode: 0644]
weblib/ricciott/prova7.ma [new file with mode: 0644]

diff --git a/weblib/cicm2012/disambiguation.ma b/weblib/cicm2012/disambiguation.ma
new file mode 100644 (file)
index 0000000..f1665c1
--- /dev/null
@@ -0,0 +1,5 @@
+include "cicm2012/naturals.ma".
+include "cicm2012/rationals.ma".
+include "cicm2012/reals.ma".
+
+lemma associative_plus : ∀x,y,z.add (add x y) z = add x (add y z).
\ No newline at end of file
diff --git a/weblib/cicm2012/naturals.ma b/weblib/cicm2012/naturals.ma
new file mode 100644 (file)
index 0000000..0098289
--- /dev/null
@@ -0,0 +1,5 @@
+include "basics/logic.ma".
+
+\ 5img class="anchor" src="icons/tick.png" id="N"\ 6axiom N : Type[0].
+\ 5img class="anchor" src="icons/tick.png" id="add"\ 6axiom add : \ 5a href="cic:/matita/cicm2012/naturals/N.dec"\ 6N\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/naturals/N.dec"\ 6N\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/naturals/N.dec"\ 6N\ 5/a\ 6.
+\ 5img class="anchor" src="icons/tick.png" id="times"\ 6axiom times : \ 5a href="cic:/matita/cicm2012/naturals/N.dec"\ 6N\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/naturals/N.dec"\ 6N\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/naturals/N.dec"\ 6N\ 5/a\ 6.
diff --git a/weblib/cicm2012/part1.ma b/weblib/cicm2012/part1.ma
new file mode 100644 (file)
index 0000000..757ed30
--- /dev/null
@@ -0,0 +1,201 @@
+(* 
+\ 5h1\ 6Matita Interactive Tutorial\ 5/h1\ 6
+\ 5b\ 6The goat, the wolf and the cabbage\ 5/b\ 6
+A farmer need to transfer a goat, a wolf and a cabbage across a river, but there
+is only one place available on his boat. Furthermore, the goat will eat the 
+cabbage if they are left alone on the same bank, and similarly the wolf will eat
+the goat. The problem consists in bringing all three items safely across the 
+river. 
+
+Our first data type defines the two banks of the river, which will be named east
+and west. It is a simple example of enumerated type, defined by explicitly 
+declaring all its elements. The type itself is called "bank".
+Before giving its definition we "include" the file "logic.ma" that contains a 
+few preliminary notions not worth discussing for the moment.
+*)
+
+include "basics/logic.ma".
+
+inductive bank: Type[0] ≝
+| east : bank 
+| west : bank.
+
+(* We can now define a simple function computing, for each bank of the river, the
+opposite one. *)
+
+definition opposite ≝ λs.
+match s with
+  [ east ⇒ west
+  | west ⇒ east
+  ].
+
+(* The goal window. Computation. *)
+lemma east_to_west : opposite east = west.
+normalize
+// qed.
+
+(* the computation step is not necessary *)
+lemma west_to_east : opposite west = east.
+// qed.
+
+(* A slightly more complex problem consists in proving that opposite is idempotent *)
+lemma idempotent_opposite : ∀x. opposite (opposite x) = x.
+#b
+cases b
+// qed.
+
+(* 
+\ 5h2 class="section"\ 6Predicates\ 5/h2\ 6
+Instead of working with functions, it is sometimes convenient to work with
+predicates. For instance, instead of defining a function computing the opposite 
+bank, we could declare a predicate stating when two banks are opposite to each 
+other. Only two cases are possible, leading naturally to the following 
+definition:
+*)
+
+inductive opp : bank → bank → Prop ≝ 
+| east_west : opp east west
+| west_east : opp west east.
+
+(* In precisely the same way as "bank" is the smallest type containing east and
+west, opp is the smallest predicate containing the two sub-cases east_west and
+weast_east. If you have some familiarity with Prolog, you may look at opp as the
+predicate defined by the two clauses - in this case, the two facts - ast_west and
+west_east.
+
+Between opp and opposite we have the following relation:
+    opp a b iff a = opposite b
+Let us prove it, starting from the left to right implication, first *)
+
+lemma opp_to_opposite: ∀a,b. opp a b → a = opposite b.
+#a #b #oppab
+cases oppab // qed.
+
+(* 
+\ 5h2 class="section"\ 6Rewriting\ 5/h2\ 6
+Let us come to the opposite direction. *)
+
+lemma opposite_to_opp: ∀a,b. a = opposite b → opp a b.
+#a #b #eqa
+>eqa
+cases b // qed.
+
+(*
+\ 5h2 class="section"\ 6Records\ 5/h2\ 6
+It is time to proceed with our formalization of the farmer's problem. 
+A state of the system is defined by the position of four items: the goat, the 
+wolf, the cabbage, and the boat. The simplest way to declare such a data type
+is to use a record.
+*)
+
+record state : Type[0] ≝
+  {goat_pos : bank;
+   wolf_pos : bank;
+   cabbage_pos: bank;
+   boat_pos : bank}.
+
+(* When you define a record named foo, the system automatically defines a record 
+constructor named mk_foo. To construct a new record you pass as arguments to 
+mk_foo the values of the record fields *)
+
+definition start ≝ mk_state east east east east.
+definition end ≝ mk_state west west west west.
+
+(* We must now define the possible moves. A natural way to do it is in the form 
+of a relation (a binary predicate) over states.
+
+          g ⊥ g1         
+ ----------------------- (move_goat) 
+ (g,w,c,g) ↦ (g1,w,c,g1) 
+
+          w ⊥ w1         
+ ----------------------- (move_wolf) 
+ (g,w,c,w) ↦ (g,w1,c,w1) 
+
+          c ⊥ c1         
+ ----------------------- (move_cabbage) 
+ (g,w,c,c) ↦ (g,w,c1,c1) 
+
+         b ⊥ b1         
+ ---------------------- (move_boat) 
+ (g,w,c,b) ↦ (g,w,c,b1) 
+
+*)
+
+inductive move : state → state → Prop ≝
+| move_goat: ∀g,g1,w,c. opp g g1 → move (mk_state g w c g) (mk_state g1 w c g1)
+  (* We can move the goat from a bank g to the opposite bank g1 if and only if the
+     boat is on the same bank g of the goat and we move the boat along with it. *)
+| move_wolf: ∀g,w,w1,c. opp w w1 → move (mk_state g w c w) (mk_state g w1 c w1)
+| move_cabbage: ∀g,w,c,c1.opp c c1 → move (mk_state g w c c) (mk_state g w c1 c1)
+| move_boat: ∀g,w,c,b,b1. opp b b1 → move (mk_state g w c b) (mk_state g w c b1).
+
+(* A state is safe if either the goat is on the same bank of the boat, or both 
+the wolf and the cabbage are on the opposite bank of the goat. *)
+
+inductive safe_state : state → Prop ≝
+| with_boat : ∀g,w,c.safe_state (mk_state g w c g)
+| opposite_side : ∀g,g1,b.opp g g1 → safe_state (mk_state g g1 g1 b).
+
+(* Finally, a state y is reachable from x if either there is a single move 
+leading from x to y, or there is a safe state z such that z is reachable from x 
+and there is a move leading from z to y *)
+
+inductive reachable : state → state → Prop ≝
+| one : ∀x,y.move x y → reachable x y
+| more : ∀x,z,y. \ 5span style="text-decoration: underline;"\ 6\ 5/span\ 6reachable x z → safe_state z → \ 5span style="text-decoration: underline;"\ 6\ 5/span\ 6move z y → reachable x y.
+
+(* 
+\ 5h2 class="section"\ 6Automation\ 5/h2\ 6
+Remarkably, Matita is now able to solve the problem by itself, provided
+you allow automation to exploit more resources. The command /n/ is similar to
+//, where n is a measure of this complexity: in particular it is a bound to
+the depth of the expected proof tree (more precisely, to the number of nested
+applicative nodes). In this case, there is a solution in six moves, and we
+need a few more applications to handle reachability, and side conditions. 
+The magic number to let automation work is, in this case, 9. *)
+
+lemma problem: reachable start end.
+normalize /9/ qed. 
+
+(* 
+\ 5h2 class="section"\ 6A complex interactive proof\ 5/h2\ 6
+Let us now try to derive the proof in a more interactive way. Of course, we
+expect to need several moves to transfer all items from a bank to the other, so 
+we should start our proof by applying "more". Matita syntax for invoking the 
+application of a property named foo is to write "@foo". In general, the philosophy 
+of Matita is to describe each proof of a property P as a structured collection of 
+objects involved in the proof, prefixed by simple modalities (#,<,@,...) explaining 
+the way it is actually used (e.g. for introduction, rewriting, in an applicative 
+step, and so on).
+
+- tinycals
+- implicit arguments
+*)
+
+lemma problem1: reachable start end.
+normalize @more                     
+(*                                  
+   - multiple conjectures           
+   - tinycals                       
+ *)                                 
+  [  @(mk_state east west west east) 
+  || //
+  | /2/ ] 
+
+(* implicit arguments *)
+@(more ? (mk_state east west west west))
+/2/ 
+
+(* vectors of implicit arguments *) 
+@(more … (move_wolf … ))
+
+(* focusing *)
+[4: @east] /2/
+
+(* Alternatively, we can directly instantiate the bank into the move. Let
+us complete the proof in this, very readable way. *)
+@(more … (move_goat west … )) /2/
+@(more … (move_cabbage ?? east … )) /2/
+@(more … (move_boat ??? west … )) /2/
+@one /2/ qed.
\ No newline at end of file
diff --git a/weblib/cicm2012/part2.ma b/weblib/cicm2012/part2.ma
new file mode 100644 (file)
index 0000000..ad8f06a
--- /dev/null
@@ -0,0 +1,314 @@
+(* 
+\ 5h1 class="section"\ 6Naif Set Theory\ 5/h1\ 6
+*)
+include "basics/types.ma".
+include "basics/bool.ma".
+(* 
+In this Chapter we shall develop a naif theory of sets represented as 
+characteristic predicates over some universe \ 5code\ 6A\ 5/code\ 6, that is as objects of type 
+A→Prop. 
+For instance the empty set is defined by the always false function: *)
+
+\ 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 "∅" non associative with precedence 90 for @{'empty_set}.
+interpretation "empty set" 'empty_set = (empty_set ?).
+
+(* Similarly, a singleton set containing an element a, is defined
+by the characteristic function asserting equality with a *)
+
+\ 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 @{'singl $x}. *)
+interpretation "singleton" 'singl x = (singleton ? x).
+
+(* The membership relation between an element of type A and a set S:A →Prop is
+simply the predicate resulting from the application of S to a.
+The operations of union, intersection, complement and substraction 
+are easily defined in terms of the propositional connectives of dijunction,
+conjunction and negation *)
+
+\ 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="difference"\ 6definition difference := λ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 "difference" 'minus a b = (difference ? a b).
+
+(* Finally, we use implication to define the inclusion relation between
+sets *)
+
+\ 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).
+
+(* 
+\ 5h2 class="section"\ 6Set Equality\ 5/h2\ 6
+Two sets are equals if and only if they have the same elements, that is,
+if the two characteristic functions are extensionally equivalent: *) 
+
+\ 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.
+notation "A =1 B" non associative with precedence 45 for @{'eqP $A $B}.
+interpretation "extensional equality" 'eqP a b = (eqP ? a b).
+
+(* the fact it defines an equivalence relation must be explicitly proved: *)
+
+\ 5img class="anchor" src="icons/tick.png" id="eqP_sym"\ 6lemma eqP_sym: ∀U.∀A,B:U →Prop. 
+  A =1 B → B =1 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 =1 B → B =1 C → A =1 C.
+#U #A #B #C #eqAB #eqBC #a @\ 5a href="cic:/matita/basics/logic/iff_trans.def(2)"\ 6iff_trans\ 5/a\ 6 // qed.
+
+(* For the same reason, we must also prove that all the operations behave well
+with respect to eqP: *)
+
+\ 5img class="anchor" src="icons/tick.png" id="eqP_union_r"\ 6lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. 
+  A =1 C  → A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B =1 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 =1 C  → A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B =1 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 =1 C  → A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B =1 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 =1 C  → A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B =1 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 =1 C  → A \ 5a title="difference" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 B =1 C \ 5a title="difference" 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 =1 C  → A \ 5a title="difference" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 B =1 A \ 5a title="difference" 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.
+
+(* 
+\ 5h2 class="section"\ 6Simple properties of sets\ 5/h2\ 6
+We can now prove several properties of the previous set-theoretic operations. 
+In particular, union is commutative and associative, and the empty set is an 
+identity element: *)
+
+\ 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 =1 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 =1 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 =1 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,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,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.   
+
+(* In the same way we prove commutativity and associativity for set 
+interesection *)
+
+\ 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 =1 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="cap_assoc"\ 6lemma cap_assoc: ∀U.∀A,B,C:U→Prop.
+  A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (B \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C) =1 (A \ 5a title="intersection" 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 % [ * #Aw * /\ 5span class="autotactic"\ 63\ 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\ 5span class="autotactic"\ 6\ 5span class="autotrace"\ 6\ 5/span\ 6\ 5/span\ 6| * * /\ 5span class="autotactic"\ 63\ 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.
+
+(* We can also easily prove idempotency for union and intersection *)
+
+\ 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 =1 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 =1 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. 
+
+(* We conclude our examples with a couple of distributivity theorems, and a 
+characterization of substraction in terms of interesection and complementation. *)
+
+\ 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 =1 (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="difference" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 C =1 (A \ 5a title="difference" 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="difference" 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="substract_def"\ 6lemma substract_def:∀U.∀A,B:U→Prop. A\ 5a title="difference" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6B =1 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.
+
+(* 
+\ 5h2 class="section"\ 6Bool vs. Prop\ 5/h2\ 6
+In several situation it is important to assume to have a decidable equality 
+between elements of a set U, namely a boolean function eqb: U→U→bool such that
+for any pair of elements a and b in U, (eqb x y) is true if and only if x=y. 
+A set equipped with such an equality is called a DeqSet: *)
+
+\ 5img class="anchor" src="icons/tick.png" id="DeqSet"\ 6record DeqSet : Type[1] ≝ { carr :> Type[0];
+   eqb: carr → carr → \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6;
+   eqb_true: ∀x,y. (eqb x y \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (x \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 y)
+}.
+
+(* We use the notation == to denote the decidable equality, to distinguish it
+from the propositional equality. In particular, a term of the form a==b is a 
+boolean, while a=b is a proposition. *)
+
+notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }.
+interpretation "eqb" 'eqb a b = (eqb ? a b).
+
+(* 
+\ 5h2 class="section"\ 6Small Scale Reflection\ 5/h2\ 6
+It is convenient to have a simple way to reflect a proof of the fact 
+that (eqb a b) is true into a proof of the proposition (a = b); to this aim, 
+we introduce two operators "\P" and "\b". *)
+
+notation "\P H" non associative with precedence 90 
+  for @{(proj1 … (eqb_true ???) $H)}. 
+
+notation "\b H" non associative with precedence 90 
+  for @{(proj2 … (eqb_true ???) $H)}. 
+  
+(* If H:eqb a b = true, then \P H: a = b, and conversely if h:a = b, then
+\b h: eqb a b = true. Let us see an example of their use: the following 
+statement asserts that we can reflect a proof that eqb a b is false into
+a proof of the proposition a ≠ b. *)
+
+\ 5img class="anchor" src="icons/tick.png" id="eqb_false"\ 6lemma eqb_false: ∀S:\ 5a href="cic:/matita/cicm2012/part2/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀a,b:S. 
+  (\ 5a href="cic:/matita/cicm2012/part2/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 ? a b) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 \ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 a \ 5a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b.
+#S #a #b % 
+(* same tactic on two goals *)
+#H 
+  [@(\ 5a href="cic:/matita/basics/logic/not_to_not.def(3)"\ 6not_to_not\ 5/a\ 6 … \ 5a href="cic:/matita/basics/bool/not_eq_true_false.def(3)"\ 6not_eq_true_false\ 5/a\ 6) #H1 
+   <H @\ 5a href="cic:/matita/basics/logic/sym_eq.def(2)"\ 6sym_eq\ 5/a\ 6 @(\b H1)
+  |cases (\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6true_or_false\ 5/a\ 6 (\ 5a href="cic:/matita/cicm2012/part2/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 ? a b)) // #H1 @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 @(\ 5a href="cic:/matita/basics/logic/absurd.def(2)"\ 6absurd\ 5/a\ 6 … (\P H1) H)
+  ]
+qed.
+(* We also introduce two operators "\Pf" and "\bf" to reflect a proof
+of (a==b)=false into a proof of a≠b, and vice-versa *) 
+
+notation "\Pf H" non associative with precedence 90 
+  for @{(proj1 … (eqb_false ???) $H)}. 
+
+notation "\bf H" non associative with precedence 90 
+  for @{(proj2 … (eqb_false ???) $H)}. 
+
+(* The following statement proves that propositional equality in a 
+DeqSet is decidable in the traditional sense, namely either a=b or a≠b *)
+
\ 5img class="anchor" src="icons/tick.png" id="dec_eq"\ 6lemma dec_eq: ∀S:\ 5a href="cic:/matita/cicm2012/part2/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀a,b:S. a \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 b \ 5a title="logical or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 a \ 5a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b.
+#S #a #b cases (\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6true_or_false\ 5/a\ 6 (\ 5a href="cic:/matita/cicm2012/part2/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 ? a b)) #H
+  [%1 @(\P H) | %2 @(\Pf H)]
+qed.
+
+(* 
+\ 5h2 class="section"\ 6Unification Hints\ 5/h2\ 6
+A simple example of a set with a decidable equality is bool. We first define 
+the boolean equality beqb, then prove that beqb b1 b2 is true if and only if 
+b1=b2, and finally build the type DeqBool by instantiating the DeqSet record 
+with the previous information *)
+
+\ 5img class="anchor" src="icons/tick.png" id="beqb"\ 6definition beqb ≝ λb1,b2.
+  match b1 with [ true ⇒ b2 | false ⇒ \ 5a href="cic:/matita/basics/bool/notb.def(1)"\ 6notb\ 5/a\ 6 b2].
+
+notation < "a == b" non associative with precedence 45 for @{beqb $a $b }.
+\ 5img class="anchor" src="icons/tick.png" id="beqb_true"\ 6lemma beqb_true: ∀b1,b2. \ 5a href="cic:/matita/cicm2012/part2/beqb.def(2)"\ 6beqb\ 5/a\ 6 b1 b2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 \ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b1 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 b2.
+#b1 #b2 cases b1 cases b2 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. 
+
+\ 5img class="anchor" src="icons/tick.png" id="DeqBool"\ 6definition DeqBool ≝ \ 5a href="cic:/matita/cicm2012/part2/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 \ 5a href="cic:/matita/cicm2012/part2/beqb.def(2)"\ 6beqb\ 5/a\ 6 \ 5a href="cic:/matita/cicm2012/part2/beqb_true.def(4)"\ 6beqb_true\ 5/a\ 6.
+
+(* At this point, we would expect to be able to prove things like the
+following: for any boolean b, if b==false is true then b=false. 
+Unfortunately, this would not work, unless we declare b of type 
+DeqBool (change the type in the following statement and see what 
+happens). *)
+
+\ 5img class="anchor" src="icons/tick.png" id="exhint"\ 6example exhint: ∀b:\ 5a href="cic:/matita/cicm2012/part2/DeqBool.def(5)"\ 6DeqBool\ 5/a\ 6. (b=\ 5a title="eqb" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 → b\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
+#b #H @(\P H) 
+qed.
+
+(* The point is that == expects in input a pair of objects whose type must be the 
+carrier of a DeqSet; bool is indeed the carrier of DeqBool, but the type inference 
+system has no knowledge of it (it is an information that has been supplied by the 
+user, and stored somewhere in the library). More explicitly, the type inference 
+inference system, would face an unification problem consisting to unify bool 
+against the carrier of something (a metavaribale) and it has no way to synthetize 
+the answer. To solve this kind of situations, matita provides a mechanism to hint 
+the system the expected solution. A unification hint is a kind of rule, whose rhd 
+is the unification problem, containing some metavariables X1, ..., Xn, and whose 
+left hand side is the solution suggested to the system, in the form of equations 
+Xi=Mi. The hint is accepted by the system if and only the solution is correct, that
+is, if it is a unifier for the given problem.
+To make an example, in the previous case, the unification problem is bool = carr X,
+and the hint is to take X= mk_DeqSet bool beqb true. The hint is correct, since 
+bool is convertible with (carr (mk_DeqSet bool beb true)). *)
+
+unification hint  0 \ 5a href="cic:/fakeuri.def(1)" title="hint_decl_Type1"\ 6\ 5/a\ 6 ; 
+    X ≟ \ 5a href="cic:/matita/cicm2012/part2/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 \ 5a href="cic:/matita/cicm2012/part2/beqb.def(2)"\ 6beqb\ 5/a\ 6 \ 5a href="cic:/matita/cicm2012/part2/beqb_true.def(4)"\ 6beqb_true\ 5/a\ 6
+(* ---------------------------------------- *) ⊢ 
+    \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 ≡ \ 5a href="cic:/matita/cicm2012/part2/carr.fix(0,0,2)"\ 6carr\ 5/a\ 6 X.
+    
+unification hint  0 \ 5a href="cic:/fakeuri.def(1)" title="hint_decl_Type0"\ 6\ 5/a\ 6 b1,b2:\ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6
+    X ≟ \ 5a href="cic:/matita/cicm2012/part2/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 \ 5a href="cic:/matita/cicm2012/part2/beqb.def(2)"\ 6beqb\ 5/a\ 6 \ 5a href="cic:/matita/cicm2012/part2/beqb_true.def(4)"\ 6beqb_true\ 5/a\ 6
+(* ---------------------------------------- *) ⊢ 
+    \ 5a href="cic:/matita/cicm2012/part2/beqb.def(2)"\ 6beqb\ 5/a\ 6 b1 b2 ≡ \ 5a href="cic:/matita/cicm2012/part2/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 X b1 b2.
+    
+(* After having provided the previous hints, we may rewrite example exhint 
+declaring b of type bool. *)
+\ 5img class="anchor" src="icons/tick.png" id="exhint1"\ 6example exhint1: ∀b:\ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6. (b =\ 5a title="eqb" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 → b \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
+#b #H @(\P H)
+qed.
+
+(* The cartesian product of two DeqSets is still a DeqSet. To prove
+this, we must as usual define the boolen equality function, and prove
+it correctly reflects propositional equality. *)
+
+\ 5img class="anchor" src="icons/tick.png" id="eq_pairs"\ 6definition eq_pairs ≝
+  λA,B:\ 5a href="cic:/matita/cicm2012/part2/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.λp1,p2:A\ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6B.(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p1 =\ 5a title="eqb" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p2) \ 5a title="boolean and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p1 =\ 5a title="eqb" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p2).
+
+\ 5img class="anchor" src="icons/tick.png" id="eq_pairs_true"\ 6lemma eq_pairs_true: ∀A,B:\ 5a href="cic:/matita/cicm2012/part2/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀p1,p2:A\ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6B.
+  \ 5a href="cic:/matita/cicm2012/part2/eq_pairs.def(4)"\ 6eq_pairs\ 5/a\ 6 A B p1 p2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 \ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 p1 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 p2.
+#A #B * #a1 #b1 * #a2 #b2 %
+  [#H cases (\ 5a href="cic:/matita/basics/bool/andb_true.def(5)"\ 6andb_true\ 5/a\ 6 …H) normalize #eqa #eqb >(\P eqa) >(\P eqb) //
+  |#H destruct normalize >(\b (\ 5a href="cic:/matita/basics/logic/eq.con(0,1,2)"\ 6refl\ 5/a\ 6 … a2)) >(\b (\ 5a href="cic:/matita/basics/logic/eq.con(0,1,2)"\ 6refl\ 5/a\ 6 … b2)) //
+  ]
+qed.
+
+\ 5img class="anchor" src="icons/tick.png" id="DeqProd"\ 6definition DeqProd ≝ λA,B:\ 5a href="cic:/matita/cicm2012/part2/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.
+  \ 5a href="cic:/matita/cicm2012/part2/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 (A\ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6B) (\ 5a href="cic:/matita/cicm2012/part2/eq_pairs.def(4)"\ 6eq_pairs\ 5/a\ 6 A B) (\ 5a href="cic:/matita/cicm2012/part2/eq_pairs_true.def(6)"\ 6eq_pairs_true\ 5/a\ 6 A B).
+
+(* Having a unification problem of the kind T1×T2 = carr X, what kind 
+of hint can we give to the system? We expect T1 to be the carrier of a
+DeqSet C1, T2 to be the carrier of a DeqSet C2, and X to be DeqProd C1 C2.
+This is expressed by the following hint: *)
+
+unification hint  0 \ 5a href="cic:/fakeuri.def(1)" title="hint_decl_Type1"\ 6\ 5/a\ 6 C1,C2; 
+    T1 ≟ \ 5a href="cic:/matita/cicm2012/part2/carr.fix(0,0,2)"\ 6carr\ 5/a\ 6 C1,
+    T2 ≟ \ 5a href="cic:/matita/cicm2012/part2/carr.fix(0,0,2)"\ 6carr\ 5/a\ 6 C2,
+    X ≟ \ 5a href="cic:/matita/cicm2012/part2/DeqProd.def(7)"\ 6DeqProd\ 5/a\ 6 C1 C2
+(* ---------------------------------------- *) ⊢ 
+    T1\ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6T2 ≡ \ 5a href="cic:/matita/cicm2012/part2/carr.fix(0,0,2)"\ 6carr\ 5/a\ 6 X.
+
+unification hint  0 \ 5a href="cic:/fakeuri.def(1)" title="hint_decl_Type0"\ 6\ 5/a\ 6 T1,T2,p1,p2; 
+    X ≟ \ 5a href="cic:/matita/cicm2012/part2/DeqProd.def(7)"\ 6DeqProd\ 5/a\ 6 T1 T2
+(* ---------------------------------------- *) ⊢ 
+    \ 5a href="cic:/matita/cicm2012/part2/eq_pairs.def(4)"\ 6eq_pairs\ 5/a\ 6 T1 T2 p1 p2 ≡ \ 5a href="cic:/matita/cicm2012/part2/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 X p1 p2.
+
+\ 5img class="anchor" src="icons/tick.png" id="hint2"\ 6example hint2: ∀b1,b2. 
+  〈b1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6=\ 5a title="eqb" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,b2\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 → 〈b1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,b2\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6.
+#b1 #b2 #H @(\P H)
+qed.
diff --git a/weblib/cicm2012/rationals.ma b/weblib/cicm2012/rationals.ma
new file mode 100644 (file)
index 0000000..cb44258
--- /dev/null
@@ -0,0 +1,5 @@
+include "basics/logic.ma".
+
+\ 5img class="anchor" src="icons/tick.png" id="Q"\ 6axiom Q : Type[0].
+\ 5img class="anchor" src="icons/tick.png" id="add"\ 6axiom add : \ 5a href="cic:/matita/cicm2012/rationals/Q.dec"\ 6Q\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/rationals/Q.dec"\ 6Q\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/rationals/Q.dec"\ 6Q\ 5/a\ 6.
+\ 5img class="anchor" src="icons/tick.png" id="times"\ 6axiom times : \ 5a href="cic:/matita/cicm2012/rationals/Q.dec"\ 6Q\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/rationals/Q.dec"\ 6Q\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/rationals/Q.dec"\ 6Q\ 5/a\ 6.
\ No newline at end of file
diff --git a/weblib/cicm2012/reals.ma b/weblib/cicm2012/reals.ma
new file mode 100644 (file)
index 0000000..4b4b431
--- /dev/null
@@ -0,0 +1,5 @@
+include "basics/logic.ma".
+
+\ 5img class="anchor" src="icons/tick.png" id="R"\ 6axiom R : Type[0].
+\ 5img class="anchor" src="icons/tick.png" id="add"\ 6axiom add : \ 5a href="cic:/matita/cicm2012/reals/R.dec"\ 6R\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/reals/R.dec"\ 6R\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/reals/R.dec"\ 6R\ 5/a\ 6.
+\ 5img class="anchor" src="icons/tick.png" id="times"\ 6axiom times : \ 5a href="cic:/matita/cicm2012/reals/R.dec"\ 6R\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/reals/R.dec"\ 6R\ 5/a\ 6 → \ 5a href="cic:/matita/cicm2012/reals/R.dec"\ 6R\ 5/a\ 6.
\ No newline at end of file
index 6745f532a52d9603b1d8beb542f7fa4411bd87cf..c830cde3be64345ff64d6d236ce79f496d4ada8e 100644 (file)
@@ -1 +1,3 @@
-axiom pluto : Prop.
+(* test *)
+
+axiom pluto : Prop.
\ No newline at end of file
diff --git a/weblib/part1.ma b/weblib/part1.ma
new file mode 100644 (file)
index 0000000..9bb135c
--- /dev/null
@@ -0,0 +1,378 @@
+(* 
+\ 5h1\ 6Matita Interactive Tutorial\ 5/h1\ 6
+This is an interactive tutorial. To let you interact on line with the system, 
+you must first of all register yourself.
+
+Before starting, let us briefly explain the meaning of the menu buttons. 
+With the Advance and Retract buttons you respectively perform and undo single 
+computational steps. Each step consists in reading a user command, and processing
+it. The part of the user input file (called script) already executed by the 
+system will be colored, and will not be editable any more. The advance bottom 
+will also automatically advance the focus of the window, but you can inspect the 
+whole file using the scroll bars, as usual. Comments are skipped.
+Try to advance and retract a few steps, to get the feeling of the system. You can 
+also come back here by using the top button, that takes you at the beginning of 
+a file. The play button is meant to process the script up to a position 
+previously selected by the user; the bottom button execute the whole script. 
+That's it: we are\ 5span style="font-family: Verdana,sans-serif;"\ 6 \ 5/span\ 6now ready to start.
+
+\ 5h2 class="section"\ 6Data types, functions and theorems\ 5/h2\ 6
+Matita is both a programming language and a theorem proving environment:
+you define datatypes and programs, and then prove properties on them.
+Very few things are built-in: not even booleans or logical connectives
+(but you may of course use libraries, as in normal programming languages). 
+The main philosophy of the system is to let you define your own data-types 
+and functions using a powerful computational mechanism based on the 
+declaration of inductive types. 
+
+Let us start this tutorial with a simple example based on the following well 
+known problem.
+
+\ 5b\ 6The goat, the wolf and the cabbage\ 5/b\ 6
+A farmer need to transfer a goat, a wolf and a cabbage across a river, but there
+is only one place available on his boat. Furthermore, the goat will eat the 
+cabbage if they are left alone on the same bank, and similarly the wolf will eat
+the goat. The problem consists in bringing all three items safely across the 
+river. 
+
+Our first data type defines the two banks of the river, which will be named east
+and west. It is a simple example of enumerated type, defined by explicitly 
+declaring all its elements. The type itself is called "bank".
+Before giving its definition we "include" the file "logic.ma" that contains a 
+few preliminary notions not worth discussing for the moment.
+*)
+
+include "basics/logic.ma".
+
+inductive bank: Type[0] ≝
+| east : bank 
+| west : bank.
+
+(* We can now define a simple function computing, for each bank of the river, the
+opposite one. *)
+
+definition opposite ≝ λs.
+match s with
+  [ east ⇒ \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6
+  | west ⇒ \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6
+  ].
+
+(* Functions are live entities, and can be actually computed. To check this, let
+us state the property that the opposite bank of east is west; every lemma needs a 
+name for further reference, and we call it "east_to_west". *)
+lemma east_to_west : \ 5a href="cic:/matita/tutorial/chapter1/opposite.def(1)"\ 6opposite\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6.
+
+(* 
+\ 5h2 class="section"\ 6The goal window\ 5/h2\ 6
+If you stop the execution here you will see a new window on the  right side
+of the screen: it is the goal window, providing a sequent like representation of
+the form
+
+B1
+B2
+....
+Bk
+-----------------------
+A
+
+for each open goal remaining to be solved. A is the conclusion of the goal and 
+B1, ..., Bk is its context, that is the set of current hypothesis and type 
+declarations. In this case, we have only one goal, and its context is empty. 
+The proof proceeds by typing commands to the system. In this case, we
+want to evaluate the function, that can be done by invoking the  "normalize"
+command:
+*)
+
+normalize
+
+(* By executing it - just type the advance command - you will see that the goal
+has changed to west = west, by reducing the subexpression (opposite east). 
+You may use the retract bottom to undo the step, if you wish. 
+
+The new goal west = west is trivial: it is just a consequence of reflexivity.
+Such trivial steps can be just closed in Matita by typing a double slash. 
+We complete the proof by the qed command, that instructs the system to store the
+lemma performing some book-keeping operations. 
+*)
+
+// qed.
+
+(* In exactly the same way, we can prove that the opposite side of west is east.
+In this case, we avoid the unnecessary simplification step: // will take care of 
+it. *) 
+
+lemma west_to_east : \ 5a href="cic:/matita/tutorial/chapter1/opposite.def(1)"\ 6opposite\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6.
+// qed.
+
+(*
+\ 5h2 class="section"\ 6Introduction\ 5/h2\ 6
+A slightly more complex problem consists in proving that opposite is idempotent *)
+
+lemma idempotent_opposite : ∀x. \ 5a href="cic:/matita/tutorial/chapter1/opposite.def(1)"\ 6opposite\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter1/opposite.def(1)"\ 6opposite\ 5/a\ 6 x) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 x.
+
+(* we start the proof moving x from the conclusion into the context, that is a 
+(backward) introduction step. Matita syntax for an introduction step is simply 
+the sharp character followed by the name of the item to be moved into the 
+context. This also allows us to rename the item, if needed: for instance if we 
+wish to rename x into b (since it is a bank), we just type #b. *)
+
+#b
+
+(* See the effect of the execution in the goal window on the right: b has been 
+added to the context (replacing x in the conclusion); moreover its implicit type 
+"bank" has been made explicit. 
+*)
+
+(*
+\ 5h2 class="section"\ 6Case analysis\ 5/h2\ 6
+But how are we supposed to proceed, now? Simplification cannot help us, since b
+is a variable: just try to call normalize and you will see that it has no effect.
+The point is that we must proceed by cases according to the possible values of b,
+namely east and west. To this aim, you must invoke the cases command, followed by
+the name of the hypothesis (more generally, an arbitrary expression) that must be
+the object of the case analysis (in our case, b).
+*)
+
+cases b
+
+(* Executing the previous command has the effect of opening two subgoals, 
+corresponding to the two cases b=east and b=west: you may switch from one to the 
+other by using the hyperlinks over the top of the goal window. 
+Both goals can be closed by trivial computations, so we may use // as usual.
+If we had to treat each subgoal in a different way, we should focus on each of 
+them in turn, in a way that will be described at the end of this section.
+*)
+
+// qed.
+
+(* 
+\ 5h2 class="section"\ 6Predicates\ 5/h2\ 6
+Instead of working with functions, it is sometimes convenient to work with
+predicates. For instance, instead of defining a function computing the opposite 
+bank, we could declare a predicate stating when two banks are opposite to each 
+other. Only two cases are possible, leading naturally to the following 
+definition:
+*)
+
+inductive opp : \ 5a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"\ 6bank\ 5/a\ 6 → \ 5a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"\ 6bank\ 5/a\ 6 → Prop ≝ 
+| east_west : opp \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6
+| west_east : opp \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6.
+
+(* In precisely the same way as "bank" is the smallest type containing east and
+west, opp is the smallest predicate containing the two sub-cases east_west and
+weast_east. If you have some familiarity with Prolog, you may look at opp as the
+predicate defined by the two clauses - in this case, the two facts - ast_west and
+west_east.
+
+Between opp and opposite we have the following relation:
+    opp a b iff a = opposite b
+Let us prove it, starting from the left to right implication, first *)
+
+lemma opp_to_opposite: ∀a,b. \ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 a b → a \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/opposite.def(1)"\ 6opposite\ 5/a\ 6 b.
+(* We start the proof introducing a, b and the hypothesis opp a b, that we
+call oppab. *)
+#a #b #oppab
+
+(* Now we proceed by cases on the possible proofs of (opp a b), that is on the 
+possible shapes of oppab. By definition, there are only two possibilities, 
+namely east_west or west_east. Both subcases are trivial, and can be closed by
+automation *)
+
+cases oppab // qed.
+
+(* 
+\ 5h2 class="section"\ 6Rewriting\ 5/h2\ 6
+Let us come to the opposite direction. *)
+
+lemma opposite_to_opp: ∀a,b. a \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/opposite.def(1)"\ 6opposite\ 5/a\ 6 b → \ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 a b.
+
+(* As usual, we start introducing a, b and the hypothesis (a = opposite b), 
+that we call eqa. *)
+
+#a #b #eqa
+
+(* The right way to proceed, now, is by rewriting a into (opposite b). We do
+this by typing ">eqa". If we wished to rewrite in the opposite direction, namely
+opposite b into a, we would have typed "<eqa". *)
+
+>eqa
+
+(* We conclude the proof by cases on b. *)
+
+cases b // qed.
+
+(*
+\ 5h2 class="section"\ 6Records\ 5/h2\ 6
+It is time to proceed with our formalization of the farmer's problem. 
+A state of the system is defined by the position of four items: the goat, the 
+wolf, the cabbage, and the boat. The simplest way to declare such a data type
+is to use a record.
+*)
+
+record state : Type[0] ≝
+  {goat_pos : \ 5a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"\ 6bank\ 5/a\ 6;
+   wolf_pos : \ 5a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"\ 6bank\ 5/a\ 6;
+   cabbage_pos: \ 5a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"\ 6bank\ 5/a\ 6;
+   boat_pos : \ 5a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"\ 6bank\ 5/a\ 6}.
+
+(* When you define a record named foo, the system automatically defines a record 
+constructor named mk_foo. To construct a new record you pass as arguments to 
+mk_foo the values of the record fields *)
+
+definition start ≝ \ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6.
+definition end ≝ \ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6.
+
+(* We must now define the possible moves. A natural way to do it is in the form 
+of a relation (a binary predicate) over states. *)
+
+inductive move : \ 5a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"\ 6state\ 5/a\ 6 → \ 5a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"\ 6state\ 5/a\ 6 → Prop ≝
+| move_goat: ∀g,g1,w,c. \ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 g g1 → move (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c g) (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g1 w c g1)
+  (* We can move the goat from a bank g to the opposite bank g1 if and only if the
+     boat is on the same bank g of the goat and we move the boat along with it. *)
+| move_wolf: ∀g,w,w1,c. \ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 w w1 → move (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c w) (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w1 c w1)
+| move_cabbage: ∀g,w,c,c1.\ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 c c1 → move (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c c) (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c1 c1)
+| move_boat: ∀g,w,c,b,b1. \ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 b b1 → move (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c b) (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c b1).
+
+(* A state is safe if either the goat is on the same bank of the boat, or both 
+the wolf and the cabbage are on the opposite bank of the goat. *)
+
+inductive safe_state : \ 5a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"\ 6state\ 5/a\ 6 → Prop ≝
+| with_boat : ∀g,w,c.safe_state (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g w c g)
+| opposite_side : ∀g,g1,b.\ 5a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"\ 6opp\ 5/a\ 6 g g1 → safe_state (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 g g1 g1 b).
+
+(* Finally, a state y is reachable from x if either there is a single move 
+leading from x to y, or there is a safe state z such that z is reachable from x 
+and there is a move leading from z to y *)
+
+inductive reachable : \ 5a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"\ 6state\ 5/a\ 6 → \ 5a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"\ 6state\ 5/a\ 6 → Prop ≝
+| one : ∀x,y.\ 5a href="cic:/matita/tutorial/chapter1/move.ind(1,0,0)"\ 6move\ 5/a\ 6 x y → reachable x y
+| more : ∀x,z,y. \ 5span style="text-decoration: underline;"\ 6\ 5/span\ 6reachable x z → \ 5a href="cic:/matita/tutorial/chapter1/safe_state.ind(1,0,0)"\ 6safe_state\ 5/a\ 6 z → \ 5span style="text-decoration: underline;"\ 6\ 5/span\ 6\ 5a href="cic:/matita/tutorial/chapter1/move.ind(1,0,0)"\ 6move\ 5/a\ 6 z y → reachable x y.
+
+(* 
+\ 5h2 class="section"\ 6Automation\ 5/h2\ 6
+Remarkably, Matita is now able to solve the problem by itslef, provided
+you allow automation to exploit more resources. The command /n/ is similar to
+//, where n is a measure of this complexity: in particular it is a bound to
+the depth of the expected proof tree (more precisely, to the number of nested
+applicative nodes). In this case, there is a solution in six moves, and we
+need a few more applications to handle reachability, and side conditions. 
+The magic number to let automation work is, in this case, 9.  *)
+
+lemma problem: \ 5a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"\ 6reachable\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/start.def(1)"\ 6start\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/end.def(1)"\ 6end\ 5/a\ 6.
+normalize /\ 5span class="autotactic"\ 69\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"\ 6one\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"\ 6with_boat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"\ 6opposite_side\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"\ 6move_goat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,2,0)"\ 6move_wolf\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"\ 6move_cabbage\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"\ 6move_boat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"\ 6east_west\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"\ 6west_east\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed. 
+
+(* 
+\ 5h2 class="section"\ 6Application\ 5/h2\ 6
+Let us now try to derive the proof in a more interactive way. Of course, we
+expect to need several moves to transfer all items from a bank to the other, so 
+we should start our proof by applying "more". Matita syntax for invoking the 
+application of a property named foo is to write "@foo". In general, the philosophy 
+of Matita is to describe each proof of a property P as a structured collection of 
+objects involved in the proof, prefixed by simple modalities (#,<,@,...) explaining 
+the way it is actually used (e.g. for introduction, rewriting, in an applicative 
+step, and so on).
+*)
+
+lemma problem1: \ 5a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"\ 6reachable\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/start.def(1)"\ 6start\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/end.def(1)"\ 6end\ 5/a\ 6.
+normalize @\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6
+
+(* 
+\ 5h2 class="section"\ 6Focusing\ 5/h2\ 6
+After performing the previous application, we have four open subgoals:
+
+  X : STATE
+  Y : reachable [east,east,east,east] X
+  W : safe_state X
+  Z : move X [west,west,west,west]
+That is, we must guess a state X, such that this is reachable from start, it is 
+safe, and there is a move leading from X to end. All goals are active, that is
+emphasized by the fact that they are all red. Any command typed by the user is
+normally applied in parallel to all active goals, but clearly we must proceed 
+here is a different way for each of them. The way to do it, is by structuring
+the script using the following syntax: [...|... |...|...] where we typically have
+as many cells inside the brackets as the number of the active subgoals. The
+interesting point is that we can associate with the three symbol "[", "|" and
+"]" a small-step semantics that allow to execute them individually. In particular
+
+- the operator "[" opens a new focusing section for the currently active goals,
+  and focus on the first of them
+- the operator "|" shift the focus to the next goal
+- the operator "]" close the focusing section, falling back to the previous level
+  and adding to it all remaining goals not yet closed
+
+Let us see the effect of the "[" on our proof:
+*)
+
+  [  
+
+(* As you see, only the first goal has now the focus on. Moreover, all goals got
+a progressive numerical label, to help designating them, if needed. 
+We can now proceed in several possible ways. The most straightforward way is to 
+provide the intermediate state, that is [east,west,west,east]. We can do it, by 
+just applying this term. *)
+
+   @(\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6
+
+(* This application closes the goal; at present, no goal has the focus on.
+In order to act on the next goal, we must focus on it using the "|" operator. In
+this case, we would like to skip the next goal, and focus on the trivial third 
+subgoal: a simple way to do it, is by retyping "|". The proof that 
+[east,west,west,east] is safe is trivial and can be done with //.*)
+
+  || //
+
+(*
+We then advance to the next goal, namely the fact that there is a move from 
+[east,west,west,east] to [west,west,west,west]; this is trivial too, but it 
+requires /2/ since move_goat opens an additional subgoal. By applying "]" we
+refocus on the skipped goal, going back to a situation similar to the one we
+started with. *)
+
+  | /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"\ 6move_goat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"\ 6east_west\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ ] 
+
+(* 
+\ 5h2 class="section"\ 6Implicit arguments\ 5/h2\ 6
+Let us perform the next step, namely moving back the boat, in a sligtly 
+different way. The more operation expects as second argument the new 
+intermediate state, hence instead of applying more we can apply this term
+already instatated on the next intermediate state. As first argument, we
+type a question mark that stands for an implicit argument to be guessed by
+the system. *)
+
+@(\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6 ? (\ 5a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"\ 6mk_state\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6))
+
+(* We now get three independent subgoals, all actives, and two of them are 
+trivial. We\ 5span style="font-family: Verdana,sans-serif;"\ 6 \ 5/span\ 6can just apply automation to all of them, and it will close the two
+trivial goals. *)
+
+/\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"\ 6opposite_side\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"\ 6move_boat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"\ 6east_west\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"\ 6west_east\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
+
+(* Let us come to the next step, that consists in moving the wolf. Suppose that 
+instead of specifying the next intermediate state, we prefer to specify the next 
+move. In the spirit of the previous example, we can do it in the following way 
+*)
+
+@(\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,2,0)"\ 6move_wolf\ 5/a\ 6 … ))
+
+(* The dots stand here for an arbitrary number of implicit arguments, to be 
+guessed by the system. 
+Unfortunately, the previous move is not enough to fully instantiate the new 
+intermediate state: a bank B remains unknown. Automation cannot help here,
+since all goals depend from this bank and automation refuses to close some
+subgoals instantiating other subgoals remaining open (the instantiation could
+be arbitrary). The simplest way to proceed is to focus on the bank, that is
+the fourth subgoal, and explicitly instatiate it. Instead of repeatedly using "|",
+we can perform focusing by typing "4:" as described by the following command. *)
+
+[4: @\ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6] /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"\ 6with_boat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"\ 6east_west\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
+
+(* Alternatively, we can directly instantiate the bank into the move. Let
+us complete the proof in this, very readable way. *)
+
+@(\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"\ 6move_goat\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 … )) /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"\ 6with_boat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"\ 6west_east\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
+@(\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"\ 6move_cabbage\ 5/a\ 6 ?? \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"\ 6east\ 5/a\ 6 … )) /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"\ 6opposite_side\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"\ 6east_west\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"\ 6west_east\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
+@(\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"\ 6more\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"\ 6move_boat\ 5/a\ 6 ??? \ 5a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"\ 6west\ 5/a\ 6 … )) /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"\ 6with_boat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"\ 6west_east\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
+@\ 5a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"\ 6one\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"\ 6move_goat\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"\ 6east_west\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
\ No newline at end of file
diff --git a/weblib/ricciott/prova7.ma b/weblib/ricciott/prova7.ma
new file mode 100644 (file)
index 0000000..5fe5dce
--- /dev/null
@@ -0,0 +1 @@
+(* new script *)
\ No newline at end of file