]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/didactic/duality.ma
8b7f5a399a6894962442facc4b630b534c8254c5
[helm.git] / helm / software / matita / contribs / didactic / duality.ma
1
2 include "nat/minus.ma".
3
4 let rec max n m on n ≝ match n - m with [ O => m | _ => n]. 
5 let rec min n m on n ≝ match n - m with [ O => n | _ => m]. 
6 definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f].
7 notation > "'if' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 90 for @{ 'if_then_else $e $t $f }.
8 notation < "'if' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 90 for @{ 'if_then_else $e $t $f }.
9 interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else _ e t f).
10
11
12 inductive Formula : Type ≝
13 | FBot: Formula
14 | FTop: Formula
15 | FAtom: nat → Formula
16 | FAnd: Formula → Formula → Formula
17 | FOr: Formula → Formula → Formula
18 | FImpl: Formula → Formula → Formula
19 | FNot: Formula → Formula
20 .
21
22 let rec sem (v: nat → nat) (F: Formula) on F ≝
23  match F with
24   [ FBot ⇒ 0
25   | FTop ⇒ 1
26   | FAtom n ⇒ min (v n) 1
27   | FAnd F1 F2 ⇒ min (sem v F1) (sem v F2)
28   | FOr F1 F2 ⇒ max (sem v F1) (sem v F2)
29   | FImpl F1 F2 ⇒ max (1 - sem v F1) (sem v F2)
30   | FNot F1 ⇒ 1 - (sem v F1)
31   ]
32 .
33
34 notation < "[[ \nbsp term 19 a \nbsp ]] \nbsp \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
35 notation > "[[ term 19 a ]] \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
36 notation > "[[ term 19 a ]]_ term 90 v" non associative with precedence 90 for @{ sem $v $a }.
37 interpretation "Semantic of Formula" 'semantics v a = (sem v a).
38
39 lemma sem_bool : ∀F,v. [[ F ]]_v = 0 ∨ [[ F ]]_v = 1.
40 intros; elim F; simplify;
41 [left;reflexivity;
42 |right;reflexivity;
43 |cases (v n);[left;|cases n1;right;]reflexivity;
44 |4,5,6: cases H; cases H1; rewrite > H2; rewrite > H3; simplify;
45         first [ left;reflexivity | right; reflexivity ].
46 |cases H; rewrite > H1; simplify;[right|left]reflexivity;]
47 qed.
48
49 lemma min_bool : ∀n. min n 1 = 0 ∨ min n 1 = 1.
50 intros; cases n; [left;reflexivity] cases n1; right; reflexivity;
51 qed.  
52
53 lemma min_max : ∀F,G,v.
54   min (1 - [[F]]_v) (1 - [[G]]_v) = 1 - max [[F]]_v [[G]]_v.
55 intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1;
56 simplify; reflexivity;
57 qed.
58
59 lemma max_min : ∀F,G,v.
60   max (1 - [[F]]_v) (1 - [[G]]_v) = 1 - min [[F]]_v [[G]]_v.
61 intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1;
62 simplify; reflexivity;
63 qed.
64
65 let rec negate (F: Formula) on F ≝
66  match F with
67   [ FBot ⇒ FBot
68   | FTop ⇒ FTop
69   | FAtom n ⇒ FNot (FAtom n)
70   | FAnd F1 F2 ⇒ FAnd (negate F1) (negate F2)
71   | FOr F1 F2 ⇒ FOr (negate F1) (negate F2)
72   | FImpl F1 F2 ⇒ FImpl (negate F1) (negate F2)
73   | FNot F ⇒ FNot (negate F)
74   ].
75
76
77 definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]_v = [[ F2 ]]_v.
78 notation "hvbox(a \nbsp break mstyle color #0000ff (≡) \nbsp b)"  non associative with precedence 45 for @{ 'equivF $a $b }.
79 notation > "a ≡ b" non associative with precedence 50 for @{ equiv $a $b }.
80 interpretation "equivalence for Formulas" 'equivF a b = (equiv a b).
81
82 lemma equiv_rewrite : ∀F1,F2,F3. F1 ≡ F2 → F1 ≡ F3 → F2 ≡ F3. intros; intro; autobatch. qed.
83
84 let rec dualize (F : Formula) on F : Formula ≝
85   match F with
86   [ FBot ⇒ FTop
87   | FTop ⇒ FBot
88   | FAtom n ⇒ FAtom n
89   | FAnd F1 F2 ⇒ FOr (dualize F1) (dualize F2)
90   | FOr F1 F2 ⇒ FAnd (dualize F1) (dualize F2)
91   | FImpl F1 F2 ⇒ FAnd (FNot (dualize F1)) (dualize F2)
92   | FNot F ⇒ FNot (dualize F)
93   ].
94
95 definition invert ≝
96  λv:ℕ -> ℕ. λx. if eqb (min (v x) 1) 0 then 1 else 0.
97  
98 (*DOCBEGIN
99
100 Il linguaggio di dimostrazione di Matita
101 ========================================
102
103 Per dimostrare questo teorema in modo agevole è necessario utilizzare il 
104 seguente comando:
105
106 *  `symmetry` 
107
108   Quando la conclusuine è `a = b` permette di cambiarla in `b = a`.
109
110 DOCEND*)
111 theorem negate_invert:
112  ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]_v=[[ F ]]_(invert v).
113 assume F:Formula.
114 assume v:(ℕ→ℕ).
115 we proceed by induction on F to prove ([[ negate F ]]_v=[[ F ]]_(invert v)).
116   case FBot .
117    the thesis becomes ([[ negate FBot ]]_v=[[ FBot ]]_(invert v)).
118   done.
119   case FTop .
120    the thesis becomes ([[ negate FTop ]]_v=[[ FTop ]]_(invert v)).
121   done.
122   case FAtom.
123    assume n : ℕ.
124   the thesis becomes ([[ negate (FAtom n) ]]_v=[[ FAtom n ]]_(invert v)).
125   the thesis becomes (1 - (min (v n) 1)= min (invert v n) 1).
126   the thesis becomes (1 - (min (v n) 1)= min (if eqb (min (v n) 1) 0 then 1 else 0) 1).
127   by min_bool we proved (min (v n) 1 = 0 ∨ min (v n) 1 = 1) (H1);
128   we proceed by cases on (H1) to prove (1 - (min (v n) 1)= min (if eqb (min (v n) 1) 0 then 1 else 0) 1).
129   case Left.
130     conclude 
131         (1 - (min (v n) 1)) 
132       = (1 - 0) by H.
133       = 1.
134     symmetry.
135     conclude 
136         (min (if eqb (min (v n) 1) O then 1 else O) 1)
137       = (min (if eqb 0 0 then 1 else O) 1) by H.
138       = (min 1 1).
139       = 1.
140   done.
141   case Right.
142     conclude 
143         (1 - (min (v n) 1)) 
144       = (1 - 1) by H.
145       = 0.
146     symmetry.
147     conclude 
148         (min (if eqb (min (v n) 1) O then 1 else O) 1)
149       = (min (if eqb 1 0 then 1 else O) 1) by H.
150       = (min 0 1).
151       = 0.
152   done.
153   case FAnd.
154     assume f : Formula.
155     by induction hypothesis we know
156       ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
157     assume f1 : Formula.
158     by induction hypothesis we know
159      ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1).
160     the thesis becomes
161      ([[ negate (FAnd f f1) ]]_v=[[ FAnd f f1 ]]_(invert v)).
162     the thesis becomes
163      (min [[ negate f ]]_v [[ negate f1]]_v = [[ FAnd f f1 ]]_(invert v)).
164     conclude 
165         (min [[ negate f ]]_v [[ negate f1]]_v)
166       = (min [[ f ]]_(invert v) [[ negate f1]]_v) by H.
167       = (min [[ f ]]_(invert v) [[ f1]]_(invert v)) by H1.
168   done.
169   case FOr.
170     assume f : Formula.
171     by induction hypothesis we know
172       ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
173     assume f1 : Formula.
174     by induction hypothesis we know
175      ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1).
176     the thesis becomes
177      ([[ negate (FOr f f1) ]]_v=[[ FOr f f1 ]]_(invert v)).
178     the thesis becomes
179      (max [[ negate f ]]_v [[ negate f1]]_v = [[ FOr f f1 ]]_(invert v)).
180     conclude 
181         (max [[ negate f ]]_v [[ negate f1]]_v)
182       = (max [[ f ]]_(invert v) [[ negate f1]]_v) by H.
183       = (max [[ f ]]_(invert v) [[ f1]]_(invert v)) by H1.
184   done.
185   case FImpl.
186     assume f : Formula.
187     by induction hypothesis we know
188       ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
189     assume f1 : Formula.
190     by induction hypothesis we know
191      ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1).
192     the thesis becomes
193      ([[ negate (FImpl f f1) ]]_v=[[ FImpl f f1 ]]_(invert v)).
194     the thesis becomes
195      (max (1 - [[ negate f ]]_v) [[ negate f1]]_v = [[ FImpl f f1 ]]_(invert v)).
196     conclude 
197         (max (1 - [[ negate f ]]_v) [[ negate f1]]_v)
198       = (max (1 - [[ f ]]_(invert v)) [[ negate f1]]_v) by H.
199       = (max (1 - [[ f ]]_(invert v)) [[ f1]]_(invert v)) by H1.
200   done.
201   case FNot.
202     assume f : Formula.
203     by induction hypothesis we know
204       ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
205    the thesis becomes
206      ([[ negate (FNot f) ]]_v=[[ FNot f ]]_(invert v)).
207    the thesis becomes
208      (1 - [[ negate f ]]_v=[[ FNot f ]]_(invert v)).
209    conclude (1 - [[ negate f ]]_v) = (1 - [[f]]_(invert v)) by H.
210  done.  
211 qed.   
212
213 (*
214 lemma negate_fun : ∀F,G. F ≡ G → negate F ≡ negate G.
215 intros; intro v; rewrite > (negate_invert ? v);rewrite > (negate_invert ? v);
216 apply H;
217 qed.
218 *)
219
220 theorem negate_fun:
221  ∀F:Formula.∀G:Formula.F ≡ G→negate F ≡ negate G.
222  assume F:Formula.
223  assume G:Formula.
224  suppose (F ≡ G) (H).
225  the thesis becomes (negate F ≡ negate G).
226  the thesis becomes (∀v:ℕ→ℕ.[[ negate F ]]_v=[[ negate G ]]_v).
227  assume v:(ℕ→ℕ).
228  conclude 
229      [[ negate F ]]_v
230    = [[ F ]]_(invert v) by negate_invert.
231    = [[ G ]]_(invert v) by H.
232    = [[ negate G ]]_v by negate_invert.
233  done.  
234 qed.
235 (*
236 lemma not_dualize_eq_negate : ∀F. FNot (dualize F) ≡ negate F.
237 intros; intro; elim F; intros; try reflexivity;
238 [1,2: simplify in ⊢ (? ? ? %); rewrite <(H); rewrite <(H1);
239         [rewrite < (min_max (dualize f) (dualize f1) v); reflexivity;
240         |rewrite < (max_min (dualize f) (dualize f1) v); reflexivity;]
241 |3: change in ⊢ (? ? ? %) with [[FImpl (negate f) (negate f1)]]_v;
242     change in ⊢ (? ? ? %) with (max (1 - [[negate f]]_v) [[negate f1]]_v);
243     rewrite <H1; rewrite <H;
244     rewrite > (max_min (FNot (dualize f)) ((dualize f1)) v);reflexivity;
245 |4: simplify; rewrite < H; reflexivity;]
246 qed.
247 *)
248
249 theorem not_dualize_eq_negate:
250  ∀F:Formula.negate F ≡ FNot (dualize F).
251  assume F:Formula.
252  the thesis becomes (∀v:ℕ→ℕ.[[negate F]]_v=[[FNot (dualize F)]]_v).
253  assume v:(ℕ→ℕ).
254  we proceed by induction on F to prove ([[negate F]]_v=[[FNot (dualize F)]]_v).
255  case FBot .
256    the thesis becomes ([[ negate FBot ]]_v=[[ FNot (dualize FBot) ]]_v).
257  done.
258  case FTop .
259    the thesis becomes ([[ negate FTop ]]_v=[[ FNot (dualize FTop) ]]_v).
260  done.
261  case FAtom.
262    assume n : ℕ.
263    the thesis becomes ([[ negate (FAtom n) ]]_v=[[ FNot (dualize (FAtom n)) ]]_v).
264  done.
265  case FAnd.
266    assume f : Formula.
267    by induction hypothesis we know
268      ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
269    assume f1 : Formula.
270    by induction hypothesis we know
271     ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1).
272    the thesis becomes
273     ([[ negate (FAnd f f1) ]]_v=[[ FNot (dualize (FAnd f f1)) ]]_v).
274    the thesis becomes
275     (min [[ negate f ]]_v [[ negate f1 ]]_v=[[ FNot (dualize (FAnd f f1)) ]]_v).
276    conclude 
277        (min [[ negate f ]]_v [[ negate f1 ]]_v)
278      = (min [[ FNot (dualize f) ]]_v [[ negate f1 ]]_v) by H.    
279      = (min [[ FNot (dualize f) ]]_v [[ FNot (dualize f1) ]]_v) by H1.
280      = (min (1 - [[ dualize f ]]_v) (1 - [[ dualize f1 ]]_v)).
281      = (1 - (max [[ dualize f ]]_v [[ dualize f1 ]]_v)) by min_max.
282  done.
283  case FOr.
284    assume f : Formula.
285    by induction hypothesis we know
286      ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
287    assume f1 : Formula.
288    by induction hypothesis we know
289     ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1).
290    the thesis becomes
291     ([[ negate (FOr f f1) ]]_v=[[ FNot (dualize (FOr f f1)) ]]_v).
292    the thesis becomes
293     (max [[ negate f ]]_v [[ negate f1 ]]_v=[[ FNot (dualize (FOr f f1)) ]]_v).
294    conclude 
295        (max [[ negate f ]]_v [[ negate f1 ]]_v)
296      = (max [[ FNot (dualize f) ]]_v [[ negate f1 ]]_v) by H.    
297      = (max [[ FNot (dualize f) ]]_v [[ FNot (dualize f1) ]]_v) by H1.
298      = (max (1 - [[ dualize f ]]_v) (1 - [[ dualize f1 ]]_v)).
299      = (1 - (min [[ dualize f ]]_v [[ dualize f1 ]]_v)) by max_min.
300  done.
301  case FImpl.
302    assume f : Formula.
303    by induction hypothesis we know
304      ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
305    assume f1 : Formula.
306    by induction hypothesis we know
307     ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1).
308    the thesis becomes
309     ([[ negate (FImpl f f1) ]]_v=[[ FNot (dualize (FImpl f f1)) ]]_v).
310    the thesis becomes
311     (max (1 - [[ negate f ]]_v) [[ negate f1 ]]_v=[[ FNot (dualize (FImpl f f1)) ]]_v).
312    conclude 
313        (max (1-[[ negate f ]]_v) [[ negate f1 ]]_v)
314      = (max (1-[[ FNot (dualize f) ]]_v) [[ negate f1 ]]_v) by H.    
315      = (max (1-[[ FNot (dualize f) ]]_v) [[ FNot (dualize f1) ]]_v) by H1.
316      = (max (1 - [[ FNot (dualize f) ]]_v) (1 - [[ dualize f1 ]]_v)).
317      = (1 - (min [[ FNot (dualize f) ]]_v [[ dualize f1 ]]_v)) by max_min.
318  done.
319  case FNot. 
320    assume f : Formula.
321    by induction hypothesis we know
322      ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
323    the thesis becomes
324       ([[ negate (FNot f) ]]_v=[[ FNot (dualize (FNot f)) ]]_v).
325    the thesis becomes
326       (1 - [[ negate f ]]_v=[[ FNot (dualize (FNot f)) ]]_v).
327    conclude (1 - [[ negate f ]]_v) = (1 - [[ FNot (dualize f) ]]_v) by H.
328  done.
329 qed.
330
331
332 (*
333 lemma not_inj : ∀F,G. FNot F ≡ FNot G → F ≡ G.
334 intros; intro v;lapply (H v) as K; 
335 change in K with (1 - [[ F ]]_v = 1 - [[ G ]]_v);
336 cases (sem_bool F v);cases (sem_bool G v); rewrite > H1; rewrite > H2;
337 try reflexivity; rewrite > H1 in K; rewrite > H2 in K; simplify in K;
338 symmetry; assumption;
339 qed.
340 *)
341
342 theorem not_inj:
343  ∀F:Formula.∀G:Formula.FNot F ≡ FNot G→F ≡ G.
344  assume F:Formula.
345  assume G:Formula.
346  suppose (FNot F ≡ FNot G) (H).
347  the thesis becomes (F ≡ G).
348  the thesis becomes (∀v:ℕ→ℕ.[[ F ]]_v=[[ G ]]_v).
349  assume v:(ℕ→ℕ).
350  by H we proved ([[ FNot F ]]_v=[[ FNot G ]]_v) (H1).
351  by sem_bool we proved ([[ F ]]_v=O∨[[ F ]]_v=1) (H2).
352  by sem_bool we proved ([[ G ]]_v=O∨[[ G ]]_v=1) (H3).
353  we proceed by cases on H2 to prove ([[ F ]]_v=[[ G ]]_v).
354  case Left.
355    we proceed by cases on H3 to prove ([[ F ]]_v=[[ G ]]_v).
356    case Left.
357    done.
358    case Right.
359      conclude 
360          ([[ F ]]_v)
361        = 0 by H4;
362        = (1 - 1).
363        = (1 - [[G]]_v) by H5.
364        = [[ FNot G ]]_v.
365        = [[ FNot F ]]_v by H1.
366        = (1 - [[F]]_v).
367        = (1 - 0) by H4.
368        = 1.
369      done.
370  case Right.
371    we proceed by cases on H3 to prove ([[ F ]]_v=[[ G ]]_v).
372    case Left.
373      conclude 
374          ([[ F ]]_v)
375        = 1 by H4;
376        = (1 - 0).
377        = (1 - [[G]]_v) by H5.
378        = [[ FNot G ]]_v.
379        = [[ FNot F ]]_v by H1.
380        = (1 - [[F]]_v).
381        = (1 - 1) by H4.
382        = 0.
383      done.
384    case Right.
385    done.
386 qed.
387
388 (*
389 theorem duality: ∀F1,F2. F1 ≡ F2 → dualize F1 ≡  dualize F2.
390 intros; apply not_inj; intro v; rewrite > (not_dualize_eq_negate ? v);
391 rewrite > (not_dualize_eq_negate ? v); apply (negate_fun ??? v); apply H;
392 qed.
393 *)
394
395
396
397 theorem duality:
398  ∀F1:Formula.∀F2:Formula.F1 ≡ F2→dualize F1 ≡ dualize F2.
399  assume F1:Formula.
400  assume F2:Formula.
401  suppose (F1 ≡ F2) (H).
402  the thesis becomes (dualize F1 ≡ dualize F2).
403  by negate_fun we proved (negate F1 ≡ negate F2) (H1).
404  by not_dualize_eq_negate, equiv_rewrite we proved (FNot (dualize F1) ≡ negate F2) (H2).
405  by not_dualize_eq_negate, equiv_rewrite we proved (FNot (dualize F1) ≡ FNot (dualize F2)) (H3).
406  by not_inj we proved (dualize F1 ≡ dualize F2) (H4).
407  done.
408 qed.