X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdidactic%2Fduality.ma;h=7fe9082e154ef754a4f01a9a9488df5a4468e7e1;hb=72f8ff3c58c5ac927f572267386f17be39d5a026;hp=4f16daf04492267f42bc18571859215e7de8fe08;hpb=7958e5cb010a06e744de3a4f859557bd31af5083;p=helm.git diff --git a/helm/software/matita/contribs/didactic/duality.ma b/helm/software/matita/contribs/didactic/duality.ma index 4f16daf04..7fe9082e1 100644 --- a/helm/software/matita/contribs/didactic/duality.ma +++ b/helm/software/matita/contribs/didactic/duality.ma @@ -1,126 +1,644 @@ - +(* Esercitazione di logica 29/10/2008. *) + +(* Esercizio 0 + =========== + + Compilare i seguenti campi: + + Nome1: ... + Cognome1: ... + Matricola1: ... + Account1: ... + + Nome2: ... + Cognome2: ... + Matricola2: ... + Account2: ... + + Prima di abbandonare la postazione: + + * salvare il file (menu `File ▹ Save as ...`) nella directory (cartella) + /public/ con nome linguaggi_Account1.ma, ad esempio Mario Rossi, il cui + account è mrossi, deve salvare il file in /public/linguaggi_mrossi.ma + + * mandatevi via email o stampate il file. Per stampare potete usare + usare l'editor gedit che offre la funzionalità di stampa +*) + +(*DOCBEGIN + +Il teorema di dualità +===================== + +Il teorema di dualizzazione dice che date due formule `F1` ed `F2`, +se le due formule sono equivalenti (`F1 ≡ F2`) allora anche le +loro dualizzate lo sono (`dualize F1 ≡ dualize F2`). + +L'ingrediente principale è la funzione di dualizzazione di una formula `F`: + + * Scambia FTop con FBot e viceversa + + * Scambia il connettivo FAnd con FOr e viceversa + + * Sostituisce il connettivo FImpl con FAnd e nega la + prima sottoformula. + + Ad esempio la formula `A → (B ∧ ⊥)` viene dualizzata in + `¬A ∧ (B ∨ ⊤)`. + +Per dimostrare il teorema di dualizzazione in modo agevole è necessario +definire altre nozioni: + +* La funzione `negate` che presa una formula `F` ne nega gli atomi. + Ad esempio la formula `(A ∨ (⊤ → B))` deve diventare `¬A ∨ (⊤ → ¬B)`. + +* La funzione `invert` permette di invertire un mondo `v`. + Ovvero, per ogni indice di atomo `i`, se `v i` restituisce + `1` allora `(invert v) i` restituisce `0` e viceversa. + +DOCEND*) + +(* ATTENZIONE + ========== + + Non modificare quanto segue +*) include "nat/minus.ma". -definition max : nat → nat → nat ≝ λa,b:nat.let rec max n m on n ≝ match n with [ O ⇒ b | S n ⇒ match m with [ O ⇒ a | S m ⇒ max n m]] in max a b. -definition min : nat → nat → nat ≝ λa,b:nat.let rec min n m on n ≝ match n with [ O ⇒ a | S n ⇒ match m with [ O ⇒ b | S m ⇒ min n m]] in min a b. - - - +definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f]. +notation > "'if' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }. +notation < "'if' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 19 for @{ 'if_then_else $e $t $f }. +interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else _ e t f). +definition max ≝ λn,m. if eqb (n - m) 0 then m else n. +definition min ≝ λn,m. if eqb (n - m) 0 then n else m. + +(* Ripasso + ======= + + Il linguaggio delle formule, dove gli atomi sono + rapperesentati da un numero naturale +*) inductive Formula : Type ≝ | FBot: Formula -| FTop: (*BEGIN*)Formula(*END*) -| FAtom: nat → Formula (* usiamo i naturali al posto delle lettere *) +| FTop: Formula +| FAtom: nat → Formula | FAnd: Formula → Formula → Formula -| FOr: (*BEGIN*)Formula → Formula → Formula(*END*) -| FImpl: (*BEGIN*)Formula → Formula → Formula(*END*) -| FNot: (*BEGIN*)Formula → Formula(*END*) +| FOr: Formula → Formula → Formula +| FImpl: Formula → Formula → Formula +| FNot: Formula → Formula . - - -let rec sem (v: nat → nat) (F: Formula) on F ≝ +(* Esercizio 1 + =========== + + Modificare la funzione `sem` scritta nella precedente + esercitazione in modo che valga solo 0 oppure 1 nel caso degli + atomi, anche nel caso il mondo `v` restituisca un numero + maggiore di 1. + + Suggerimento: non è necessario usare il costrutto if_then_else + e tantomento il predicato di maggiore o uguale. +*) +let rec sem (v: nat → nat) (F: Formula) on F : nat ≝ match F with [ FBot ⇒ 0 - | FTop ⇒ (*BEGIN*)1(*END*) - (*BEGIN*) - | FAtom n ⇒ v n - (*END*) - | FAnd F1 F2 ⇒ (*BEGIN*)min (sem v F1) (sem v F2)(*END*) - (*BEGIN*) + | FTop ⇒ 1 + | FAtom n ⇒ (*BEGIN*)min (v n) 1(*END*) + | FAnd F1 F2 ⇒ min (sem v F1) (sem v F2) | FOr F1 F2 ⇒ max (sem v F1) (sem v F2) | FImpl F1 F2 ⇒ max (1 - sem v F1) (sem v F2) - (*END*) | FNot F1 ⇒ 1 - (sem v F1) ] . - -definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f]. -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 }. -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 }. -interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else _ e t f). +(* ATTENZIONE + ========== + + Non modificare quanto segue. +*) notation < "[[ \nbsp term 19 a \nbsp ]] \nbsp \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }. notation > "[[ term 19 a ]] \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }. notation > "[[ term 19 a ]]_ term 90 v" non associative with precedence 90 for @{ sem $v $a }. interpretation "Semantic of Formula" 'semantics v a = (sem v a). - -let rec subst (F: Formula) on F ≝ +definition v20 ≝ λx. + if eqb x 0 then 2 + else if eqb x 1 then 1 + else 0. + +(* Test 1 + ====== + + La semantica della formula `(A ∨ C)` nel mondo `v20` in cui + `A` vale `2` e `C` vale `0` deve valere `1`. + +*) +eval normalize on [[For (FAtom 0) (FAtom 2)]]_v20. + +(*DOCBEGIN + +La libreria di Matita +===================== + +Gli strumenti per la dimostrazione assistita sono corredati da +librerie di teoremi già dimostrati. Per portare a termine l'esercitazione +sono necessari i seguenti lemmi: + +* lemma `sem_bool` : `∀F,v. [[ F ]]_v = 0 ∨ [[ F ]]_v = 1` +* lemma `min_bool` : `∀n. min n 1 = 0 ∨ min n 1 = 1` +* lemma `min_max` : `∀F,G,v.min (1 - [[F]]_v) (1 - [[G]]_v) = 1 - max [[F]]_v [[G]]_v` +* lemma `max_min` : `∀F,G,v.max (1 - [[F]]_v) (1 - [[G]]_v) = 1 - min [[F]]_v [[G]]_v` +* lemma `equiv_rewrite` : `∀F1,F2,F3. F1 ≡ F2 → F1 ≡ F3 → F2 ≡ F3` + +DOCEND*) + +(* ATTENZIONE + ========== + + Non modificare quanto segue. +*) +lemma sem_bool : ∀F,v. [[ F ]]_v = 0 ∨ [[ F ]]_v = 1. intros; elim F; simplify; [left;reflexivity; |right;reflexivity; |cases (v n);[left;|cases n1;right;]reflexivity; |4,5,6: cases H; cases H1; rewrite > H2; rewrite > H3; simplify; first [ left;reflexivity | right; reflexivity ]. |cases H; rewrite > H1; simplify;[right|left]reflexivity;] qed. +lemma min_bool : ∀n. min n 1 = 0 ∨ min n 1 = 1. intros; cases n; [left;reflexivity] cases n1; right; reflexivity; qed. +lemma min_max : ∀F,G,v. min (1 - [[F]]_v) (1 - [[G]]_v) = 1 - max [[F]]_v [[G]]_v. intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1; simplify; reflexivity; qed. +lemma max_min : ∀F,G,v. max (1 - [[F]]_v) (1 - [[G]]_v) = 1 - min [[F]]_v [[G]]_v. intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1; simplify; reflexivity; qed. + +(* Esercizio 2 + =========== + + Definire per ricorsione strutturale la funzione `negate` + che presa una formula `F` ne nega gli atomi. + + Ad esempio la formula `(A ∨ (⊤ → B))` deve diventare + `¬A ∨ (⊤ → ¬B)`. +*) +let rec negate (F: Formula) on F : Formula ≝ match F with - [ FBot ⇒ FBot + [ (*BEGIN*)FBot ⇒ FBot | FTop ⇒ FTop | FAtom n ⇒ FNot (FAtom n) - | FAnd F1 F2 ⇒ FAnd (subst F1) (subst F2) - | FOr F1 F2 ⇒ FOr (subst F1) (subst F2) - | FImpl F1 F2 ⇒ FImpl (subst F1) (subst F2) - | FNot F ⇒ FNot (subst F) + | FAnd F1 F2 ⇒ FAnd (negate F1) (negate F2) + | FOr F1 F2 ⇒ FOr (negate F1) (negate F2) + | FImpl F1 F2 ⇒ FImpl (negate F1) (negate F2) + | FNot F ⇒ FNot (negate F)(*END*) ]. - +(* Test 2 + ====== + + Testare la funzione `negate`. Il risultato atteso è: + + FOr (FNot (FAtom O)) (FImpl FTop (FNot (FAtom 1))) +*) + +eval normalize on (negate (FOr (FAtom 0) (FImpl FTop (FAtom 1)))). + +(* ATTENZIONE + ========== + + Non modificare quanto segue +*) definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]_v = [[ F2 ]]_v. notation "hvbox(a \nbsp break mstyle color #0000ff (≡) \nbsp b)" non associative with precedence 45 for @{ 'equivF $a $b }. notation > "a ≡ b" non associative with precedence 50 for @{ equiv $a $b }. interpretation "equivalence for Formulas" 'equivF a b = (equiv a b). - +lemma equiv_rewrite : ∀F1,F2,F3. F1 ≡ F2 → F1 ≡ F3 → F2 ≡ F3. intros; intro; autobatch. qed. + +(* Esercizio 3 + =========== + + Definire per ricorsione strutturale la funzione di + dualizzazione di una formula `F`. Tale funzione: + + * Scambia FTop con FBot e viceversa + + * Scambia il connettivo FAnd con FOr e viceversa + + * Sostituisce il connettivo FImpl con FAnd e nega la + prima sottoformula. + + Ad esempio la formula `A → (B ∧ ⊥)` viene dualizzata in + `¬A ∧ (B ∨ ⊤)`. +*) let rec dualize (F : Formula) on F : Formula ≝ match F with - [ FBot ⇒ FTop + [ (*BEGIN*)FBot ⇒ FTop | FTop ⇒ FBot | FAtom n ⇒ FAtom n | FAnd F1 F2 ⇒ FOr (dualize F1) (dualize F2) | FOr F1 F2 ⇒ FAnd (dualize F1) (dualize F2) | FImpl F1 F2 ⇒ FAnd (FNot (dualize F1)) (dualize F2) - | FNot F ⇒ FNot (dualize F) + | FNot F ⇒ FNot (dualize F)(*END*) ]. +(* Test 3 + ====== + + Testare la funzione `dualize`. Il risultato atteso è: + + FAnd (FNot (FAtom O)) (FOr (FAtom 1) FTop) +*) + +eval normalize on (dualize (FImpl (FAtom 0) (FAnd (FAtom 1) FBot))). + +(* Spiegazione + =========== + + La funzione `invert` permette di invertire un mondo `v`. + Ovvero, per ogni indice di atomo `i`, se `v i` restituisce + `1` allora `(invert v) i` restituisce `0` e viceversa. + +*) +definition invert ≝ + λv:ℕ -> ℕ. λx. if eqb (min (v x) 1) 0 then 1 else 0. + +interpretation "Inversione del mondo" 'invert v = (invert v). + +(*DOCBEGIN + +Il linguaggio di dimostrazione di Matita +======================================== + +Per dimostrare il lemma `negate_invert` in modo agevole è necessario +utilizzare il seguente comando: + +* `symmetry` + + Quando la conclusuine è `a = b` permette di cambiarla in `b = a`. + +* by H1, H2 we proved P (H) + + Il comando `by ... we proved` visto nella scorsa esercitazione + permette di utilizzare più ipotesi o lemmi alla volta + separandoli con una virgola. + +DOCEND*) + +(* Esercizio 4 + =========== + + Dimostrare il lemma `negate_invert` che asserisce che + la semantica in un mondo `v` associato alla formula + negata di `F` e uguale alla semantica associata + a `F` in un mondo invertito. +*) +lemma negate_invert: + ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]_v=[[ F ]]_(invert v). +assume F:Formula. +assume v:(ℕ→ℕ). +we proceed by induction on F to prove ([[ negate F ]]_v=[[ F ]]_(invert v)). + case FBot. + (*BEGIN*) + the thesis becomes ([[ negate FBot ]]_v=[[ FBot ]]_(invert v)). + (*END*) + done. + case FTop. + (*BEGIN*) + the thesis becomes ([[ negate FTop ]]_v=[[ FTop ]]_(invert v)). + (*END*) + done. + case FAtom. + assume n : ℕ. + the thesis becomes ([[ negate (FAtom n) ]]_v=[[ FAtom n ]]_(invert v)). + the thesis becomes (1 - (min (v n) 1)= min (invert v n) 1). + the thesis becomes (1 - (min (v n) 1)= min (if eqb (min (v n) 1) 0 then 1 else 0) 1). + by min_bool we proved ((*BEGIN*)min (v n) 1 = 0 ∨ min (v n) 1 = 1(*END*)) (H1); + 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). + case Left. + conclude + (1 - (min (v n) 1)) + = (1 - 0) by H. + = 1. + symmetry. + conclude + (min (if eqb (min (v n) 1) O then 1 else O) 1) + = (min (if eqb 0 0 then 1 else O) 1) by H. + = (min 1 1). + = 1. + done. + case Right. + (*BEGIN*) + conclude + (1 - (min (v n) 1)) + = (1 - 1) by H. + = 0. + symmetry. + conclude + (min (if eqb (min (v n) 1) O then 1 else O) 1) + = (min (if eqb 1 0 then 1 else O) 1) by H. + = (min 0 1). + = 0. + (*END*) + done. + case FAnd. + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ f ]]_(invert v)) (H). + assume f1 : Formula. + by induction hypothesis we know + ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1). + the thesis becomes + ([[ negate (FAnd f f1) ]]_v=[[ FAnd f f1 ]]_(invert v)). + the thesis becomes + (min [[ negate f ]]_v [[ negate f1]]_v = [[ FAnd f f1 ]]_(invert v)). + conclude + (min [[ negate f ]]_v [[ negate f1]]_v) + = (min [[ f ]]_(invert v) [[ negate f1]]_v) by H. + = (min [[ f ]]_(invert v) [[ f1]]_(invert v)) by (*BEGIN*)H1(*END*). + done. + case FOr. + (*BEGIN*) + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ f ]]_(invert v)) (H). + assume f1 : Formula. + by induction hypothesis we know + ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1). + the thesis becomes + ([[ negate (FOr f f1) ]]_v=[[ FOr f f1 ]]_(invert v)). + the thesis becomes + (max [[ negate f ]]_v [[ negate f1]]_v = [[ FOr f f1 ]]_(invert v)). + conclude + (max [[ negate f ]]_v [[ negate f1]]_v) + = (max [[ f ]]_(invert v) [[ negate f1]]_v) by H. + = (max [[ f ]]_(invert v) [[ f1]]_(invert v)) by H1. + (*END*) + done. + case FImpl. + (*BEGIN*) + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ f ]]_(invert v)) (H). + assume f1 : Formula. + by induction hypothesis we know + ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1). + the thesis becomes + ([[ negate (FImpl f f1) ]]_v=[[ FImpl f f1 ]]_(invert v)). + the thesis becomes + (max (1 - [[ negate f ]]_v) [[ negate f1]]_v = [[ FImpl f f1 ]]_(invert v)). + conclude + (max (1 - [[ negate f ]]_v) [[ negate f1]]_v) + = (max (1 - [[ f ]]_(invert v)) [[ negate f1]]_v) by H. + = (max (1 - [[ f ]]_(invert v)) [[ f1]]_(invert v)) by H1. + (*END*) + done. + case FNot. + (*BEGIN*) + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ f ]]_(invert v)) (H). + the thesis becomes + ([[ negate (FNot f) ]]_v=[[ FNot f ]]_(invert v)). + the thesis becomes + (1 - [[ negate f ]]_v=[[ FNot f ]]_(invert v)). + conclude (1 - [[ negate f ]]_v) = (1 - [[f]]_(invert v)) by H. + (*END*) + done. +qed. + +(* Esercizio 5 + =========== + + Dimostrare che la funzione negate rispetta l'equivalenza. +*) +lemma negate_fun: + ∀F:Formula.∀G:Formula.F ≡ G→negate F ≡ negate G. + (*BEGIN*) + assume F:Formula. + assume G:Formula. + suppose (F ≡ G) (H). + the thesis becomes (negate F ≡ negate G). + the thesis becomes (∀v:ℕ→ℕ.[[ negate F ]]_v=[[ negate G ]]_v). + (*END*) + assume v:(ℕ→ℕ). + conclude + [[ negate F ]]_v + = [[ F ]]_(invert v) by negate_invert. + = [[ G ]]_((*BEGIN*)invert v(*BEGIN*)) by (*BEGIN*)H(*BEGIN*). + = [[ negate G ]]_(*BEGIN*)v(*BEGIN*) by (*BEGIN*)negate_invert(*END*). + done. +qed. + +(* Esercizio 6 + =========== + + Dimostrare che per ogni formula `F`, `(negae F)` equivale a + dualizzarla e negarla. +*) +lemma not_dualize_eq_negate: + ∀F:Formula.negate F ≡ FNot (dualize F). + (*BEGIN*) + assume F:Formula. + the thesis becomes (∀v:ℕ→ℕ.[[negate F]]_v=[[FNot (dualize F)]]_v). + (*END*) + assume v:(ℕ→ℕ). + we proceed by induction on F to prove ([[negate F]]_v=[[FNot (dualize F)]]_v). + case FBot. + (*BEGIN*) + the thesis becomes ([[ negate FBot ]]_v=[[ FNot (dualize FBot) ]]_v). + (*END*) + done. + case FTop. + (*BEGIN*) + the thesis becomes ([[ negate FTop ]]_v=[[ FNot (dualize FTop) ]]_v). + (*END*) + done. + case FAtom. + (*BEGIN*) + assume n : ℕ. + the thesis becomes ([[ negate (FAtom n) ]]_v=[[ FNot (dualize (FAtom n)) ]]_v). + (*END*) + done. + case FAnd. + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H). + assume f1 : Formula. + by induction hypothesis we know + ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1). + the thesis becomes + ([[ negate (FAnd f f1) ]]_v=[[ FNot (dualize (FAnd f f1)) ]]_v). + the thesis becomes + (min [[ negate f ]]_v [[ negate f1 ]]_v=[[ FNot (dualize (FAnd f f1)) ]]_v). + conclude + (min [[ negate f ]]_v [[ negate f1 ]]_v) + = (min [[ FNot (dualize f) ]]_v [[ negate f1 ]]_v) by H. + = (min [[ FNot (dualize f) ]]_v [[ FNot (dualize f1) ]]_v) by H1. + = (min (1 - [[ dualize f ]]_v) (1 - [[ dualize f1 ]]_v)). + = (1 - (max [[ dualize f ]]_v [[ dualize f1 ]]_v)) by min_max. + done. + case FOr. + (*BEGIN*) + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H). + assume f1 : Formula. + by induction hypothesis we know + ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1). + the thesis becomes + ([[ negate (FOr f f1) ]]_v=[[ FNot (dualize (FOr f f1)) ]]_v). + the thesis becomes + (max [[ negate f ]]_v [[ negate f1 ]]_v=[[ FNot (dualize (FOr f f1)) ]]_v). + conclude + (max [[ negate f ]]_v [[ negate f1 ]]_v) + = (max [[ FNot (dualize f) ]]_v [[ negate f1 ]]_v) by H. + = (max [[ FNot (dualize f) ]]_v [[ FNot (dualize f1) ]]_v) by H1. + = (max (1 - [[ dualize f ]]_v) (1 - [[ dualize f1 ]]_v)). + = (1 - (min [[ dualize f ]]_v [[ dualize f1 ]]_v)) by max_min. + (*END*) + done. + case FImpl. + (*BEGIN*) + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H). + assume f1 : Formula. + by induction hypothesis we know + ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1). + the thesis becomes + ([[ negate (FImpl f f1) ]]_v=[[ FNot (dualize (FImpl f f1)) ]]_v). + the thesis becomes + (max (1 - [[ negate f ]]_v) [[ negate f1 ]]_v=[[ FNot (dualize (FImpl f f1)) ]]_v). + conclude + (max (1-[[ negate f ]]_v) [[ negate f1 ]]_v) + = (max (1-[[ FNot (dualize f) ]]_v) [[ negate f1 ]]_v) by H. + = (max (1-[[ FNot (dualize f) ]]_v) [[ FNot (dualize f1) ]]_v) by H1. + = (max (1 - [[ FNot (dualize f) ]]_v) (1 - [[ dualize f1 ]]_v)). + = (1 - (min [[ FNot (dualize f) ]]_v [[ dualize f1 ]]_v)) by max_min. + (*END*) + done. + case FNot. + (*BEGIN*) + assume f : Formula. + by induction hypothesis we know + ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H). + the thesis becomes + ([[ negate (FNot f) ]]_v=[[ FNot (dualize (FNot f)) ]]_v). + the thesis becomes + (1 - [[ negate f ]]_v=[[ FNot (dualize (FNot f)) ]]_v). + conclude (1 - [[ negate f ]]_v) = (1 - [[ FNot (dualize f) ]]_v) by H. + (*END*) + done. +qed. + +(* Esercizio 7 + =========== + + Dimostrare che la negazione è iniettiva +*) +theorem not_inj: + ∀F,G:Formula.FNot F ≡ FNot G→F ≡ G. + (*BEGIN*) + assume F:Formula. + assume G:Formula. + suppose (FNot F ≡ FNot G) (H). + the thesis becomes (F ≡ G). + the thesis becomes (∀v:ℕ→ℕ.[[ F ]]_v=[[ G ]]_v). + (*END*) + assume v:(ℕ→ℕ). + by H we proved ([[ FNot F ]]_v=[[ FNot G ]]_v) (H1). + by sem_bool we proved ([[ F ]]_v=O ∨ [[ F ]]_v=1) (H2). + by (*BEGIN*)sem_bool(*END*) we proved ([[ G ]]_v=(*BEGIN*)O ∨ [[ G ]]_v=1(*END*)) (H3). + we proceed by cases on H2 to prove ([[ F ]]_v=[[ G ]]_v). + case Left. + we proceed by cases on H3 to prove ([[ F ]]_v=[[ G ]]_v). + case Left. + done. + case Right. + conclude + ([[ F ]]_v) + = 0 by H4; + = (1 - 1). + = (1 - [[G]]_v) by H5. + = [[ FNot G ]]_v. + = [[ FNot F ]]_v by H1. + = (1 - [[F]]_v). + = (1 - 0) by H4. + = 1. + done. + case Right. + (*BEGIN*) + we proceed by cases on H3 to prove ([[ F ]]_v=[[ G ]]_v). + case Left. + conclude + ([[ F ]]_v) + = 1 by H4; + = (1 - 0). + = (1 - [[G]]_v) by H5. + = [[ FNot G ]]_v. + = [[ FNot F ]]_v by H1. + = (1 - [[F]]_v). + = (1 - 1) by H4. + = 0. + done. + case Right. + (*END*) + done. +qed. + +(*DOCBEGIN + +La prova del teorema di dualità +=============================== + +Il teorema di dualità accennato a lezione dice che se due formule +`F1` ed `F2` sono equivalenti, allora anche le formule duali lo sono. + + ∀F1,F2:Formula. F1 ≡ F2 → dualize F1 ≡ dualize F2. + +Per dimostrare tale teorema è bene suddividere la prova in lemmi intermedi + +1. lemma `negate_invert`, dimostrato per induzione su F, utilizzando + `min_bool` + + ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]_v=[[ F ]]_(invert v). + +2. lemma `negate_fun`, conseguenza di `negate_invert` + + ∀F,G:Formula. F ≡ G → negate F ≡ negate G. + +2. lemma `not_dualize_eq_negate`, dimostrato per induzione su F, + utilizzando `max_min` e `min_max` + + ∀F:Formula. negate F ≡ FNot (dualize F) + +4. lemma `not_inj`, conseguenza di `sem_bool` + + ∀F,G:Formula. FNot F ≡ FNot G → F ≡ G -lemma : F[G/x] ≡ F - - - - - - - - - - - - - - -lemma max_n_O : ∀n. max n O = n. intros; cases n; reflexivity; qed. - -lemma min_n_n : ∀n. min n n = n. intros; elim n; [reflexivity] simplify;assumption; +Una volta dimostrati tali lemmi la prova del teorema di dualità +procede come di seguito: -lemma min_minus_n_m : ∀n,m. min (n-m) n = n-m. intro; elim n; try reflexivity; -cases m; [ -simplify in ⊢ (? ? (? % ?) ?);simplify; rewrite > (H n1); +1. Assume l'ipotesi -lemma min_one : - ∀x,y.1 - max x y = min (1 -x) (1-y). -intros; apply (nat_elim2 ???? y x); intros; -[ rewrite > max_n_O; -whd in ⊢ (? ? ? (? ? %)); + F1 ≡ F2 +2. Utilizza `negate_fun` per ottenere + negate F1 ≡ negate F2 +3. Utilizzando due volte il lemma `not_dualize_eq_negate` e il lemma + `equiv_rewrite` ottiene - + FNot (dualize F1) ≡ FNot (dualize F2) -axiom dualize2 : ∀F. FNot (dualize F) ≡ subst F. -[[ subst f ]]v = [[ f ]]_(1-v) -f=g -> subst f = subst g -not f = not g => f = g +4. Conclude utilizzando il lemma `not_inj` per ottenere la tesi -theorem dualize1: ∀F1,F2. F1 ≡ F2 → dualize F1 ≡ dualize F2. -intros; -cut (FNot F1 ≡ FNot F2); -. -cut (dualize (subst F1) ≡ dualize (subst F2)). -cut (subst (subst F1) ≡ subst (subst F2)). + dualize F1 ≡ dualize F2 -dual f = dual (subst subst f) - = (dual subst) (subst f) - = not (subst f) - = +DOCEND*) +(* Esercizio 8 + =========== + + Dimostrare il teorema di dualità +*) +theorem duality: ∀F1,F2:Formula.F1 ≡ F2 → dualize F1 ≡ dualize F2. + assume F1:Formula. + assume F2:Formula. + suppose (F1 ≡ F2) (H). + the thesis becomes (dualize F1 ≡ dualize F2). + by (*BEGIN*)negate_fun(*END*) we proved (negate F1 ≡ negate F2) (H1). + by (*BEGIN*)not_dualize_eq_negate(*END*), (*BEGIN*)equiv_rewrite(*END*), H1 we proved (FNot (dualize F1) ≡ negate F2) (H2). + by (*BEGIN*)not_dualize_eq_negate(*END*), (*BEGIN*)equiv_rewrite(*END*), H2 we proved (FNot (dualize F1) ≡ FNot (dualize F2)) (H3). + by (*BEGIN*)not_inj(*END*), H3 we proved (dualize F1 ≡ dualize F2) (H4). + done. +qed.