X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=weblib%2Ftutorial%2Fchapter1.ma;h=4cf145d60f53f72c70b08846444700cd9b1afc7e;hb=f38fd769279794d0ca73c8945eac30e8b42e59be;hp=7a961d4d2e548357719c8969cba310478f1f5750;hpb=2979a29642a4b0b16834ec69177813a6c6051e7e;p=helm.git diff --git a/weblib/tutorial/chapter1.ma b/weblib/tutorial/chapter1.ma index 7a961d4d2..4cf145d60 100644 --- a/weblib/tutorial/chapter1.ma +++ b/weblib/tutorial/chapter1.ma @@ -44,14 +44,14 @@ few preliminary notions not worth discussing for the moment. include "basics/logic.ma". -inductive bank: Type[0] ≝ +img class="anchor" src="icons/tick.png" id="bank"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. +img class="anchor" src="icons/tick.png" id="opposite"definition opposite ≝ λs. match s with [ east ⇒ a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a | west ⇒ a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a @@ -61,7 +61,7 @@ match s with 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 : a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a. +img class="anchor" src="icons/tick.png" id="east_to_west"lemma east_to_west : a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a. (* h2 class="section"The goal window/h2 @@ -102,14 +102,14 @@ lemma performing some book-keeping operations. In this case, we avoid the unnecessary simplification step: // will take care of it. *) -lemma west_to_east : a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. +img class="anchor" src="icons/tick.png" id="west_to_east"lemma west_to_east : a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. // qed. (* h2 class="section"Introduction/h2 A slightly more complex problem consists in proving that opposite is idempotent *) -lemma idempotent_opposite : ∀x. a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a (a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a x) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. +img class="anchor" src="icons/tick.png" id="idempotent_opposite"lemma idempotent_opposite : ∀x. a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a (a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a x) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a 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 @@ -155,7 +155,7 @@ other. Only two cases are possible, leading naturally to the following definition: *) -inductive opp : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a → a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a → Prop ≝ +img class="anchor" src="icons/tick.png" id="opp"inductive opp : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a → a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a → Prop ≝ | east_west : opp a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a | west_east : opp a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. @@ -169,7 +169,7 @@ 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. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a a b → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a b. +img class="anchor" src="icons/tick.png" id="opp_to_opposite"lemma opp_to_opposite: ∀a,b. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a a b → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a b. (* We start the proof introducing a, b and the hypothesis opp a b, that we call oppab. *) @@ -186,7 +186,7 @@ cases oppab // qed. h2 class="section"Rewriting/h2 Let us come to the opposite direction. *) -lemma opposite_to_opp: ∀a,b. a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a b → a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a a b. +img class="anchor" src="icons/tick.png" id="opposite_to_opp"lemma opposite_to_opp: ∀a,b. a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a b → a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a a b. (* As usual, we start introducing a, b and the hypothesis (a = opposite b), that we call eqa. *) @@ -206,12 +206,12 @@ cases b // qed. (* h2 class="section"Records/h2 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 item: the goat, the +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] ≝ +img class="anchor" src="icons/tick.png" id="state"record state : Type[0] ≝ {goat_pos : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a; wolf_pos : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a; cabbage_pos: a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a; @@ -221,13 +221,13 @@ record state : Type[0] ≝ constructor named mk_foo. To construct a new record you pass as arguments to mk_foo the values of the record fields *) -definition start ≝ a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. -definition end ≝ a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a. +img class="anchor" src="icons/tick.png" id="start"definition start ≝ a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. +img class="anchor" src="icons/tick.png" id="end"definition end ≝ a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a. (* 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 : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ +img class="anchor" src="icons/tick.png" id="move"inductive move : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ | move_goat: ∀g,g1,w,c. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a g g1 → move (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c g) (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a 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. *) @@ -238,7 +238,7 @@ inductive move : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state (* 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 : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ +img class="anchor" src="icons/tick.png" id="safe_state"inductive safe_state : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ | with_boat : ∀g,w,c.safe_state (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c g) | opposite_side : ∀g,g1,b.a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a g g1 → safe_state (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g g1 g1 b). @@ -246,7 +246,7 @@ inductive safe_state : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)" 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 : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ +img class="anchor" src="icons/tick.png" id="reachable"inductive reachable : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ | one : ∀x,y.a href="cic:/matita/tutorial/chapter1/move.ind(1,0,0)"move/a x y → reachable x y | more : ∀x,z,y. span style="text-decoration: underline;"/spanreachable x z → a href="cic:/matita/tutorial/chapter1/safe_state.ind(1,0,0)"safe_state/a z → span style="text-decoration: underline;"/spana href="cic:/matita/tutorial/chapter1/move.ind(1,0,0)"move/a z y → reachable x y. @@ -260,7 +260,7 @@ 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: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. +img class="anchor" src="icons/tick.png" id="problem"lemma problem: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. normalize /span class="autotactic"9span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"one/a, a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a, a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/move.con(0,2,0)"move_wolf/a, a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"move_cabbage/a, a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ qed. (* @@ -275,7 +275,7 @@ the way it is actually used (e.g. for introduction, rewriting, in an applicative step, and so on). *) -lemma problem1: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. +img class="anchor" src="icons/tick.png" id="problem1"lemma problem1: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. normalize @a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a (*