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 min_1_sem: ∀F,v.min 1 [[ F ]]_v = [[ F ]]_v. intros; cases (sem_bool F v); rewrite > H; reflexivity; qed.
+lemma max_0_sem: ∀F,v.max [[ F ]]_v 0 = [[ F ]]_v. intros; cases (sem_bool F v); rewrite > H; reflexivity; qed.
(*DOCBEGIN
La libreria di Matita
Per portare a termine l'esercitazione sono necessari i seguenti lemmi:
* lemma `decidable_eq_nat` : `∀x,y.x = y ∨ x ≠ y`
-* lemma `eqb_n_n` : `∀x.eqb x x = true`
* lemma `not_eq_to_eqb_false` : `∀x,y.x ≠ y → eqb x y = false`
+* lemma `eq_to_eqb_true` : `∀x,y.x = y → eqb x y = true`
+* lemma `min_1_sem` : `∀F,v.min 1 [[ F ]]_v = [[ F ]]_v`
+* lemma `max_0_sem` : `∀F,v.max [[ F ]]_v 0 = [[ F ]]_v`
Il teorema di espansione di Shannon
===================================
DOCEND*)
+definition IFTE := λA,B,C:Formula. FOr (FAnd A B) (FAnd (FNot A) C).
+
+lemma shannon_false:
+ ∀F,x,v. [[ FAtom x ]]_v = 0 → [[ F[FBot/x] ]]_v = [[ F ]]_v.
+assume F : Formula.
+assume x : ℕ.
+assume v : (ℕ → ℕ).
+suppose ([[ FAtom x ]]_v = 0) (H).
+we proceed by induction on F to prove ([[ F[FBot/x] ]]_v = [[ F ]]_v).
+case FBot.
+ the thesis becomes ([[ FBot[FBot/x] ]]_v = [[ FBot ]]_v).
+ the thesis becomes ([[ FBot ]]_v = [[ FBot ]]_v).
+ done.
+case FTop.
+ the thesis becomes ([[ FTop[FBot/x] ]]_v = [[ FTop ]]_v).
+ the thesis becomes ([[ FTop ]]_v = [[ FTop ]]_v).
+ done.
+case FAtom.
+ assume n : ℕ.
+ the thesis becomes ([[ (FAtom n)[FBot/x] ]]_v = [[ FAtom n ]]_v).
+ the thesis becomes ([[ if eqb n x then FBot else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ by decidable_eq_nat we proved (n = x ∨ n ≠ x) (H1).
+ we proceed by cases on H1 to prove ([[ if eqb n x then FBot else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ case Left.
+ by H2, eq_to_eqb_true we proved (eqb n x = true) (H3).
+ conclude
+ ([[ if eqb n x then FBot else (FAtom n) ]]_v)
+ = ([[ if true then FBot else (FAtom n) ]]_v) by H3.
+ = ([[ FBot ]]_v).
+ = 0.
+ = ([[ FAtom x ]]_v) by H.
+ = ([[ FAtom n ]]_v) by H2.
+ done.
+ case Right.
+ by H2, not_eq_to_eqb_false we proved (eqb n x = false) (H3).
+ conclude
+ ([[ if eqb n x then FBot else (FAtom n) ]]_v)
+ = ([[ if false then FBot else (FAtom n) ]]_v) by H3.
+ = ([[ FAtom n ]]_v).
+ done.
+case FAnd.
+ assume f1 : Formula.
+ by induction hypothesis we know ([[ f1[FBot/x] ]]_v = [[ f1 ]]_v) (H1).
+ assume f2 : Formula.
+ by induction hypothesis we know ([[ f2[FBot/x] ]]_v = [[ f2 ]]_v) (H2).
+ the thesis becomes ([[ (FAnd f1 f2)[FBot/x] ]]_v = [[ FAnd f1 f2 ]]_v).
+ conclude
+ ([[ (FAnd f1 f2)[FBot/x] ]]_v)
+ = ([[ FAnd (f1[FBot/x]) (f2[FBot/x]) ]]_v).
+ = (min [[ f1[FBot/x] ]]_v [[ f2[FBot/x] ]]_v).
+ = (min [[ f1 ]]_v [[ f2[FBot/x] ]]_v) by H1.
+ = (min [[ f1 ]]_v [[ f2 ]]_v) by H2.
+ = ([[ FAnd f1 f2 ]]_v).
+ done.
+case FOr.
+ assume f1 : Formula.
+ by induction hypothesis we know ([[ f1[FBot/x] ]]_v = [[ f1 ]]_v) (H1).
+ assume f2 : Formula.
+ by induction hypothesis we know ([[ f2[FBot/x] ]]_v = [[ f2 ]]_v) (H2).
+ the thesis becomes ([[ (FOr f1 f2)[FBot/x] ]]_v = [[ FOr f1 f2 ]]_v).
+ conclude
+ ([[ (FOr f1 f2)[FBot/x] ]]_v)
+ = ([[ FOr (f1[FBot/x]) (f2[FBot/x]) ]]_v).
+ = (max [[ f1[FBot/x] ]]_v [[ f2[FBot/x] ]]_v).
+ = (max [[ f1 ]]_v [[ f2[FBot/x] ]]_v) by H1.
+ = (max [[ f1 ]]_v [[ f2 ]]_v) by H2.
+ = ([[ FOr f1 f2 ]]_v).
+ done.
+case FImpl.
+ assume f1 : Formula.
+ by induction hypothesis we know ([[ f1[FBot/x] ]]_v = [[ f1 ]]_v) (H1).
+ assume f2 : Formula.
+ by induction hypothesis we know ([[ f2[FBot/x] ]]_v = [[ f2 ]]_v) (H2).
+ the thesis becomes ([[ (FImpl f1 f2)[FBot/x] ]]_v = [[ FImpl f1 f2 ]]_v).
+ conclude
+ ([[ (FImpl f1 f2)[FBot/x] ]]_v)
+ = ([[ FImpl (f1[FBot/x]) (f2[FBot/x]) ]]_v).
+ = (max (1 - [[ f1[FBot/x] ]]_v) [[ f2[FBot/x] ]]_v).
+ = (max (1 - [[ f1 ]]_v) [[ f2[FBot/x] ]]_v) by H1.
+ = (max (1 - [[ f1 ]]_v) [[ f2 ]]_v) by H2.
+ = ([[ FImpl f1 f2 ]]_v).
+ done.
+case FNot.
+ assume f : Formula.
+ by induction hypothesis we know ([[ f[FBot/x] ]]_v = [[ f ]]_v) (H1).
+ the thesis becomes ([[ (FNot f)[FBot/x] ]]_v = [[ FNot f ]]_v).
+ conclude
+ ([[ (FNot f)[FBot/x] ]]_v)
+ = ([[ FNot (f[FBot/x]) ]]_v).
+ = (1 - [[ f[FBot/x] ]]_v).
+ = (1 - [[ f ]]_v) by H1.
+ = ([[ FNot f ]]_v).
+ done.
+qed.
+
+lemma shannon_true:
+ ∀F,x,v. [[ FAtom x ]]_v = 1 → [[ F[FTop/x] ]]_v = [[ F ]]_v.
+assume F : Formula.
+assume x : ℕ.
+assume v : (ℕ → ℕ).
+suppose ([[ FAtom x ]]_v = 1) (H).
+we proceed by induction on F to prove ([[ F[FTop/x] ]]_v = [[ F ]]_v).
+case FBot.
+ the thesis becomes ([[ FBot[FTop/x] ]]_v = [[ FBot ]]_v).
+ the thesis becomes ([[ FBot ]]_v = [[ FBot ]]_v).
+ done.
+case FTop.
+ the thesis becomes ([[ FTop[FTop/x] ]]_v = [[ FTop ]]_v).
+ the thesis becomes ([[ FTop ]]_v = [[ FTop ]]_v).
+ done.
+case FAtom.
+ assume n : ℕ.
+ the thesis becomes ([[ (FAtom n)[FTop/x] ]]_v = [[ FAtom n ]]_v).
+ the thesis becomes ([[ if eqb n x then FTop else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ by decidable_eq_nat we proved (n = x ∨ n ≠ x) (H1).
+ we proceed by cases on H1 to prove ([[ if eqb n x then FTop else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ case Left.
+ by H2, eq_to_eqb_true we proved (eqb n x = true) (H3).
+ conclude
+ ([[ if eqb n x then FTop else (FAtom n) ]]_v)
+ = ([[ if true then FTop else (FAtom n) ]]_v) by H3.
+ = ([[ FTop ]]_v).
+ = 1.
+ = ([[ FAtom x ]]_v) by H.
+ = ([[ FAtom n ]]_v) by H2.
+ done.
+ case Right.
+ by H2, not_eq_to_eqb_false we proved (eqb n x = false) (H3).
+ conclude
+ ([[ if eqb n x then FTop else (FAtom n) ]]_v)
+ = ([[ if false then FTop else (FAtom n) ]]_v) by H3.
+ = ([[ FAtom n ]]_v).
+ done.
+case FAnd.
+ assume f1 : Formula.
+ by induction hypothesis we know ([[ f1[FTop/x] ]]_v = [[ f1 ]]_v) (H1).
+ assume f2 : Formula.
+ by induction hypothesis we know ([[ f2[FTop/x] ]]_v = [[ f2 ]]_v) (H2).
+ the thesis becomes ([[ (FAnd f1 f2)[FTop/x] ]]_v = [[ FAnd f1 f2 ]]_v).
+ conclude
+ ([[ (FAnd f1 f2)[FTop/x] ]]_v)
+ = ([[ FAnd (f1[FTop/x]) (f2[FTop/x]) ]]_v).
+ = (min [[ f1[FTop/x] ]]_v [[ f2[FTop/x] ]]_v).
+ = (min [[ f1 ]]_v [[ f2[FTop/x] ]]_v) by H1.
+ = (min [[ f1 ]]_v [[ f2 ]]_v) by H2.
+ = ([[ FAnd f1 f2 ]]_v).
+ done.
+case FOr.
+ assume f1 : Formula.
+ by induction hypothesis we know ([[ f1[FTop/x] ]]_v = [[ f1 ]]_v) (H1).
+ assume f2 : Formula.
+ by induction hypothesis we know ([[ f2[FTop/x] ]]_v = [[ f2 ]]_v) (H2).
+ the thesis becomes ([[ (FOr f1 f2)[FTop/x] ]]_v = [[ FOr f1 f2 ]]_v).
+ conclude
+ ([[ (FOr f1 f2)[FTop/x] ]]_v)
+ = ([[ FOr (f1[FTop/x]) (f2[FTop/x]) ]]_v).
+ = (max [[ f1[FTop/x] ]]_v [[ f2[FTop/x] ]]_v).
+ = (max [[ f1 ]]_v [[ f2[FTop/x] ]]_v) by H1.
+ = (max [[ f1 ]]_v [[ f2 ]]_v) by H2.
+ = ([[ FOr f1 f2 ]]_v).
+ done.
+case FImpl.
+ assume f1 : Formula.
+ by induction hypothesis we know ([[ f1[FTop/x] ]]_v = [[ f1 ]]_v) (H1).
+ assume f2 : Formula.
+ by induction hypothesis we know ([[ f2[FTop/x] ]]_v = [[ f2 ]]_v) (H2).
+ the thesis becomes ([[ (FImpl f1 f2)[FTop/x] ]]_v = [[ FImpl f1 f2 ]]_v).
+ conclude
+ ([[ (FImpl f1 f2)[FTop/x] ]]_v)
+ = ([[ FImpl (f1[FTop/x]) (f2[FTop/x]) ]]_v).
+ = (max (1 - [[ f1[FTop/x] ]]_v) [[ f2[FTop/x] ]]_v).
+ = (max (1 - [[ f1 ]]_v) [[ f2[FTop/x] ]]_v) by H1.
+ = (max (1 - [[ f1 ]]_v) [[ f2 ]]_v) by H2.
+ = ([[ FImpl f1 f2 ]]_v).
+ done.
+case FNot.
+ assume f : Formula.
+ by induction hypothesis we know ([[ f[FTop/x] ]]_v = [[ f ]]_v) (H1).
+ the thesis becomes ([[ (FNot f)[FTop/x] ]]_v = [[ FNot f ]]_v).
+ conclude
+ ([[ (FNot f)[FTop/x] ]]_v)
+ = ([[ FNot (f[FTop/x]) ]]_v).
+ = (1 - [[ f[FTop/x] ]]_v).
+ = (1 - [[ f ]]_v) by H1.
+ = ([[ FNot f ]]_v).
+ done.
+qed.
+
theorem shannon :
- ∀F,x,v. [[ if eqb [[FAtom x]]_v 0 then F[FBot/x] else (F[FTop/x]) ]]_v = [[F]]_v.
+ ∀F,x. IFTE (FAtom x) (F[FTop/x]) (F[FBot/x]) ≡ F.
assume F : Formula.
assume x : ℕ.
assume v : (ℕ → ℕ).
-we proceed by induction on F to prove ([[ if eqb [[FAtom x]]_v 0 then F[FBot/x] else (F[FTop/x]) ]]_v = [[F]]_v).
+the thesis becomes ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v = [[ F ]]_v).
+by sem_bool we proved ([[ FAtom x]]_v = 0 ∨ [[ FAtom x]]_v = 1) (H).
+we proceed by cases on H to prove ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v = [[ F ]]_v).
+case Left.
+ conclude
+ ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v)
+ = ([[ FOr (FAnd (FAtom x) (F[FTop/x])) (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
+ = (max [[ (FAnd (FAtom x) (F[FTop/x])) ]]_v [[ (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
+ = (max (min [[ FAtom x ]]_v [[ F[FTop/x] ]]_v) (min (1 - [[ FAtom x ]]_v) [[ F[FBot/x] ]]_v)).
+ = (max (min 0 [[ F[FTop/x] ]]_v) (min (1 - 0) [[ F[FBot/x] ]]_v)) by H.
+ = (max 0 (min 1 [[ F[FBot/x] ]]_v)).
+ = (max 0 [[ F[FBot/x] ]]_v) by min_1_sem.
+ = ([[ F[FBot/x] ]]_v).
+ = ([[ F ]]_v) by H1, shannon_false.
+ done.
+case Right.
+ conclude
+ ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v)
+ = ([[ FOr (FAnd (FAtom x) (F[FTop/x])) (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
+ = (max [[ (FAnd (FAtom x) (F[FTop/x])) ]]_v [[ (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
+ = (max (min [[ FAtom x ]]_v [[ F[FTop/x] ]]_v) (min (1 - [[ FAtom x ]]_v) [[ F[FBot/x] ]]_v)).
+ = (max (min 1 [[ F[FTop/x] ]]_v) (min (1 - 1) [[ F[FBot/x] ]]_v)) by H.
+ = (max (min 1 [[ F[FTop/x] ]]_v) (min 0 [[ F[FBot/x] ]]_v)).
+ = (max [[ F[FTop/x] ]]_v (min 0 [[ F[FBot/x] ]]_v)) by min_1_sem.
+ = (max [[ F[FTop/x] ]]_v 0).
+ = ([[ F[FTop/x] ]]_v) by max_0_sem.
+ = ([[ F ]]_v) by H1, shannon_true.
+ done.
+qed.
+
(*DOCBEGIN
-La dimostrazione
-================
+Note generali
+=============
-La dimostrazione procede per induzione sulla formula `F`. Si ricorda che:
+Si ricorda che:
1. Ogni volta che nella finestra di destra compare un simbolo `∀` oppure un
simbolo `→` è opportuno usare il comando `assume` oppure `suppose` (che
avrà tante ipotesi induttive quante sono le sue sottoformule e tali
ipotesi sono necessarie per portare a termine la dimostrazione.
-DOCEND*)
-case FBot.
-(*DOCBEGIN
-
-Il caso FBot
-------------
-
-La tesi è
-
- [[ if eqb [[FAtom x]]_v 0 then FBot[FBot/x] else (FBot[FTop/x]) ]]_v = [[FBot]]_v
-
-Si procede per casi su `eqb [[FAtom x]]_v 0`. In entrambi i casi basta
-espandere piano piano le definizioni.
-
-### I comandi da utilizzare
-
-* `the thesis becomes (...).`
-
- Afferma quale sia la tesi da dimostrare. Se ripetuto
- permette di espandere le definizioni.
-
-* `we proceed by cases on (...) to prove (...).`
-
- Permette di andare per casi su una ipotesi (quando essa è della forma
- `A ∨ B`) oppure su una espressione come `eqb n m`.
-
- Esempio: `we proceed by cases on H to prove Q.`
-
- Esempio: `we proceed by cases on (eqb x 0) to prove Q.`
-
-* `case ... .`
-
- Nelle dimostrazioni per casi o per induzioni si utulizza tale comando
- per inizia la sotto prova relativa a un caso. Esempio: `case Fbot.`
-
-* `done.`
-
- Ogni caso di una dimostrazione deve essere terminato con il comando
- `done.`
-
-DOCEND*)
- the thesis becomes ([[ if eqb [[FAtom x]]_v 0 then FBot[FBot/x] else (FBot[FTop/x]) ]]_v = [[FBot]]_v).
- we proceed by cases on (eqb [[ FAtom x ]]_v 0)
- to prove ([[ if eqb [[FAtom x]]_v 0 then FBot[FBot/x] else (FBot[FTop/x]) ]]_v = [[FBot]]_v).
- case true.
- the thesis becomes ([[ if true then FBot[FBot/x] else (FBot[FTop/x]) ]]_v = [[FBot]]_v).
- the thesis becomes ([[ FBot[FBot/x]]]_v = [[FBot]]_v).
- the thesis becomes ([[ FBot ]]_v = [[FBot]]_v).
- the thesis becomes (0 = 0).
- done.
- case false.
- done.
-case FTop.
-(*DOCBEGIN
-
-Il caso FTop
-------------
+La dimostrazione
+================
-Analogo al caso FBot
-
-DOCEND*)
- we proceed by cases on (eqb [[ FAtom x ]]_v 0)
- to prove ([[ if eqb [[FAtom x]]_v 0 then FTop[FBot/x] else (FTop[FTop/x]) ]]_v = [[FTop]]_v).
- case true.
- done.
- case false.
- done.
-case FAtom.
-(*DOCBEGIN
+...
Il caso (FAtom n)
-----------------
Analogo al caso precedente.
-### I comandi da usare
+I comandi da utilizzare
+=======================
+
+* `the thesis becomes (...).`
+
+ Afferma quale sia la tesi da dimostrare. Se ripetuto
+ permette di espandere le definizioni.
+
+* `we proceed by cases on (...) to prove (...).`
+
+ Permette di andare per casi su una ipotesi (quando essa è della forma
+ `A ∨ B`) oppure su una espressione come `eqb n m`.
+
+ Esempio: `we proceed by cases on H to prove Q.`
+
+ Esempio: `we proceed by cases on (eqb x 0) to prove Q.`
+
+* `case ... .`
+
+ Nelle dimostrazioni per casi o per induzioni si utulizza tale comando
+ per inizia la sotto prova relativa a un caso. Esempio: `case Fbot.`
+
+* `done.`
+
+ Ogni caso di una dimostrazione deve essere terminato con il comando
+ `done.`
* `assume ... : (...) .`
tesi.
DOCEND*)
- assume n : ℕ.
- the thesis becomes ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v = [[ FAtom n ]]_ v).
- by decidable_eq_nat we proved (n = x ∨ n ≠ x) (H).
- by sem_bool we proved ([[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1) (H1).
- we proceed by cases on H to prove
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v = [[ FAtom n ]]_ v).
- case Left. (* H2 : n = x *)
- we proceed by cases on H1 to prove
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v = [[ FAtom n ]]_ v).
- case Left. (* H3 : [[ FAtom x ]]_v = 0 *)
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v)
- = ([[ if eqb 0 0 then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v) by H3.
- = ([[ if true then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v).
- = ([[ (FAtom n)[ FBot/x ] ]]_v).
- = ([[ if eqb n x then FBot else (FAtom n) ]]_v).
- = ([[ if eqb n n then FBot else (FAtom n) ]]_v) by H2.
- = ([[ if true then FBot else (FAtom n) ]]_v) by eqb_n_n.
- = ([[ FBot ]]_v).
- = 0.
- = [[ FAtom x ]]_v by H3.
- = [[ FAtom n ]]_v by H2.
- done.
- case Right.
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v)
- = ([[ if eqb 1 0 then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v) by H3.
- = ([[ if false then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v).
- = ([[ (FAtom n)[ FTop/x ] ]]_v).
- = ([[ if eqb n x then FTop else (FAtom n) ]]_v).
- = ([[ if eqb n n then FTop else (FAtom n) ]]_v) by H2.
- = ([[ if true then FTop else (FAtom n) ]]_v) by eqb_n_n.
- = ([[ FTop ]]_v).
- = 1.
- = [[ FAtom x ]]_v by H3.
- = [[ FAtom n ]]_v by H2.
- done.
- case Right.
- by not_eq_to_eqb_false, H2 we proved (eqb n x = false) (H3).
- we proceed by cases on H1 to prove
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v = [[ FAtom n ]]_ v).
- case Left.
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v)
- = ([[ if eqb 0 O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v) by H4.
- = [[ (FAtom n)[ FBot/x ] ]]_v.
- = [[ if eqb n x then FBot else (FAtom n) ]]_v.
- = [[ if false then FBot else (FAtom n) ]]_v by H3.
- = [[ FAtom n ]]_v.
- done.
- case Right.
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v)
- = ([[ if eqb 1 O then (FAtom n)[ FBot/x ] else (FAtom n[ FTop/x ]) ]]_v) by H4.
- = [[ FAtom n[ FTop/x ] ]]_v.
- = [[ if eqb n x then FTop else (FAtom n) ]]_v.
- = [[ if false then FTop else (FAtom n) ]]_v by H3.
- = [[ FAtom n ]]_v.
- done.
-case FAnd.
-(*DOCBEGIN
-
-Il caso FAnd
-------------
-
-Una volta assunte eventuali sottoformule (che chiameremo f ed f1) e
-relative ipotesi induttive
-la tesi diventa `([[ if eqb [[ FAtom x ]]_v O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v = [[ FAnd f f1 ]]_v)`.
-
-Utilizzando il lemma `sem_bool` si ottiene l'ipotesi aggiuntiva
-`([[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1)`. Si procede poi per casi
-su di essa.
-
-1. caso in cui vale `[[ FAtom x ]]_v = 0`.
-
- Componendo le ipotesi induttive con `[[ FAtom x ]]_v = 0` e
- espandendo alcune definizioni si ottengono
- `([[ f[FBot/x ] ]]_v = [[ f ]]_v)` e
- `([[ f1[FBot/x ] ]]_v = [[ f1 ]]_v)`.
-
- La sotto prova termina con una catena di uguaglianze che
- lavora sul lato sinistro della tesi.
- Espandendo alcune definizioni, utilizzando
- `[[ FAtom x ]]_v = 0` e le nuove ipotesi appena ottenute
- si arriva a `(min [[ f ]]_v [[ f1 ]]_v)`.
- Tale espressione è uguale alla parte destra della conclusione.
-
-1. caso in cui vale `[[ FAtom x ]]_v = 1`.
-
- Analogo al precedente.
-
-DOCEND*)
- assume f : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H).
- assume f1 : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H1).
- the thesis becomes
- ([[ if eqb [[ FAtom x ]]_v O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v = [[ FAnd f f1 ]]_v).
- by sem_bool we proved ([[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1) (H2).
- we proceed by cases on H2 to prove
- ([[ if eqb [[ FAtom x ]]_v O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v = [[ FAnd f f1 ]]_v).
- case Left.
- by H3, H we proved
- ([[ if eqb 0 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FBot/x ] ]]_v = [[ f ]]_v) (H5).
- by H3, H1 we proved
- ([[ if eqb 0 O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H6).
- by H6 we proved ([[ f1[FBot/x ] ]]_v = [[ f1 ]]_v) (H7).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v)
- = ([[ if eqb 0 O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v) by H3.
- = ([[ if true then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v).
- = ([[ (FAnd f f1)[ FBot/x ] ]]_v).
- = ([[ FAnd (f[ FBot/x ]) (f1[ FBot/x ]) ]]_v).
- = (min [[ f[ FBot/x ] ]]_v [[ f1[ FBot/x ] ]]_v).
- = (min [[ f ]]_v [[ f1[ FBot/x ] ]]_v) by H5.
- = (min [[ f ]]_v [[ f1 ]]_v) by H6.
- = ([[ FAnd f f1 ]]_v).
- done.
- case Right.
- by H3, H we proved
- ([[ if eqb 1 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FTop/x ] ]]_v = [[ f ]]_v) (H5).
- by H3, H1 we proved
- ([[ if eqb 1 O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H6).
- by H6 we proved ([[ f1[FTop/x ] ]]_v = [[ f1 ]]_v) (H7).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v)
- = ([[ if eqb 1 O then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v) by H3.
- = ([[ if false then ((FAnd f f1)[ FBot/x ]) else ((FAnd f f1)[ FTop/x ]) ]]_v).
- = ([[ (FAnd f f1)[ FTop/x ] ]]_v).
- = ([[ FAnd (f[ FTop/x ]) (f1[ FTop/x ]) ]]_v).
- = (min [[ f[ FTop/x ] ]]_v [[ f1[ FTop/x ] ]]_v).
- = (min [[ f ]]_v [[ f1[ FTop/x ] ]]_v) by H5.
- = (min [[ f ]]_v [[ f1 ]]_v) by H6.
- = ([[ FAnd f f1 ]]_v).
- done.
-case FOr.
-(*DOCBEGIN
-
-Il caso FOr
------------
-
-Una volta assunte eventuali sottoformule e ipotesi induttive
-la tesi diventa `([[ if eqb [[ FAtom x ]]_v O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v = [[ FOr f f1 ]]_v)`.
-
-Analogo al caso FAnd.
-
-DOCEND*)
- assume f : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H).
- assume f1 : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H1).
- the thesis becomes
- ([[ if eqb [[ FAtom x ]]_v O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v = [[ FOr f f1 ]]_v).
- by sem_bool we proved ([[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1) (H2).
- we proceed by cases on H2 to prove
- ([[ if eqb [[ FAtom x ]]_v O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v = [[ FOr f f1 ]]_v).
- case Left.
- by H3, H we proved
- ([[ if eqb 0 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FBot/x ] ]]_v = [[ f ]]_v) (H5).
- by H3, H1 we proved
- ([[ if eqb 0 O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H6).
- by H6 we proved ([[ f1[FBot/x ] ]]_v = [[ f1 ]]_v) (H7).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v)
- = ([[ if eqb 0 O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v) by H3.
- = ([[ if true then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v).
- = ([[ (FOr f f1)[ FBot/x ] ]]_v).
- = ([[ FOr (f[ FBot/x ]) (f1[ FBot/x ]) ]]_v).
- = (max [[ f[ FBot/x ] ]]_v [[ f1[ FBot/x ] ]]_v).
- = (max [[ f ]]_v [[ f1[ FBot/x ] ]]_v) by H5.
- = (max [[ f ]]_v [[ f1 ]]_v) by H6.
- = ([[ FOr f f1 ]]_v).
- done.
- case Right.
- by H3, H we proved
- ([[ if eqb 1 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FTop/x ] ]]_v = [[ f ]]_v) (H5).
- by H3, H1 we proved
- ([[ if eqb 1 O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H6).
- by H6 we proved ([[ f1[FTop/x ] ]]_v = [[ f1 ]]_v) (H7).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v)
- = ([[ if eqb 1 O then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v) by H3.
- = ([[ if false then ((FOr f f1)[ FBot/x ]) else ((FOr f f1)[ FTop/x ]) ]]_v).
- = ([[ (FOr f f1)[ FTop/x ] ]]_v).
- = ([[ FOr (f[ FTop/x ]) (f1[ FTop/x ]) ]]_v).
- = (max [[ f[ FTop/x ] ]]_v [[ f1[ FTop/x ] ]]_v).
- = (max [[ f ]]_v [[ f1[ FTop/x ] ]]_v) by H5.
- = (max [[ f ]]_v [[ f1 ]]_v) by H6.
- = ([[ FOr f f1 ]]_v).
- done.
-case FImpl.
-(*DOCBEGIN
-
-Il caso FImpl
--------------
-
-Una volta assunte eventuali sottoformule e ipotesi induttive
-la tesi diventa `([[ if eqb [[ FAtom x ]]_v O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v = [[ FImpl f f1 ]]_v)`.
-
-Analogo al caso FAnd.
-
-DOCEND*)
- assume f : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H).
- assume f1 : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H1).
- the thesis becomes
- ([[ if eqb [[ FAtom x ]]_v O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v = [[ FImpl f f1 ]]_v).
- by sem_bool we proved ([[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1) (H2).
- we proceed by cases on H2 to prove
- ([[ if eqb [[ FAtom x ]]_v O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v = [[ FImpl f f1 ]]_v).
- case Left.
- by H3, H we proved
- ([[ if eqb 0 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FBot/x ] ]]_v = [[ f ]]_v) (H5).
- by H3, H1 we proved
- ([[ if eqb 0 O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H6).
- by H6 we proved ([[ f1[FBot/x ] ]]_v = [[ f1 ]]_v) (H7).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v)
- = ([[ if eqb 0 O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v) by H3.
- = ([[ if true then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v).
- = ([[ (FImpl f f1)[ FBot/x ] ]]_v).
- = ([[ FImpl (f[ FBot/x ]) (f1[ FBot/x ]) ]]_v).
- = (max (1 - [[ f[ FBot/x ] ]]_v) [[ f1[ FBot/x ] ]]_v).
- = (max (1 - [[ f ]]_v) [[ f1[ FBot/x ] ]]_v) by H5.
- = (max (1 - [[ f ]]_v) [[ f1 ]]_v) by H6.
- = ([[ FImpl f f1 ]]_v).
- done.
- case Right.
- by H3, H we proved
- ([[ if eqb 1 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FTop/x ] ]]_v = [[ f ]]_v) (H5).
- by H3, H1 we proved
- ([[ if eqb 1 O then f1[ FBot/x ] else (f1[ FTop/x ]) ]]_v = [[ f1 ]]_v) (H6).
- by H6 we proved ([[ f1[FTop/x ] ]]_v = [[ f1 ]]_v) (H7).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v)
- = ([[ if eqb 1 O then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v) by H3.
- = ([[ if false then ((FImpl f f1)[ FBot/x ]) else ((FImpl f f1)[ FTop/x ]) ]]_v).
- = ([[ (FImpl f f1)[ FTop/x ] ]]_v).
- = ([[ FImpl (f[ FTop/x ]) (f1[ FTop/x ]) ]]_v).
- = (max (1 - [[ f[ FTop/x ] ]]_v) [[ f1[ FTop/x ] ]]_v).
- = (max (1 - [[ f ]]_v) [[ f1[ FTop/x ] ]]_v) by H5.
- = (max (1 - [[ f ]]_v) [[ f1 ]]_v) by H6.
- = ([[ FImpl f f1 ]]_v).
- done.
-case FNot.
-(*DOCBEGIN
-
-Il caso FNot
-------------
-
-Una volta assunte eventuali sottoformule e ipotesi induttive
-la tesi diventa `([[ if eqb [[ FAtom x ]]_v O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v = [[ FNot f ]]_v)`.
-
-Analogo al caso FAnd.
-
-DOCEND*)
- assume f : Formula.
- by induction hypothesis we know ([[ if eqb [[ FAtom x ]]_v O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H).
- the thesis becomes
- ([[ if eqb [[ FAtom x ]]_v O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v = [[ FNot f ]]_v).
- by sem_bool we proved ([[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1) (H2).
- we proceed by cases on H2 to prove
- ([[ if eqb [[ FAtom x ]]_v O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v = [[ FNot f ]]_v).
- case Left.
- by H1, H we proved
- ([[ if eqb 0 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FBot/x ] ]]_v = [[ f ]]_v) (H5).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v)
- = ([[ if eqb 0 O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v) by H1.
- = ([[ if true then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v).
- = ([[ (FNot f)[ FBot/x ] ]]_v).
- = ([[ FNot (f[ FBot/x ]) ]]_v).
- = (1 - [[ f[ FBot/x ] ]]_v).
- = (1 - [[ f ]]_v) by H5.
- = ([[ FNot f ]]_v).
- done.
- case Right.
- by H1, H we proved
- ([[ if eqb 1 O then f[ FBot/x ] else (f[ FTop/x ]) ]]_v = [[ f ]]_v) (H4).
- by H4 we proved ([[ f[FTop/x ] ]]_v = [[ f ]]_v) (H5).
- conclude
- ([[ if eqb [[ FAtom x ]]_v O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v)
- = ([[ if eqb 1 O then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v) by H1.
- = ([[ if false then ((FNot f)[ FBot/x ]) else ((FNot f)[ FTop/x ]) ]]_v).
- = ([[ (FNot f)[ FTop/x ] ]]_v).
- = ([[ FNot (f[ FTop/x ]) ]]_v).
- = (1 - [[ f[ FTop/x ] ]]_v).
- = (1 - [[ f ]]_v) by H5.
- = ([[ FNot f ]]_v).
- done.
-qed.
-