include "nat/minus.ma".
-let rec max n m on n ≝
- match n with [ O ⇒ m | S n ⇒
- match m with [ O ⇒ S n | S m ⇒ S (max n m)]].
-let rec min n m on n ≝
- match n with [ O ⇒ O | S n ⇒
- match m with [ O ⇒ O | S m ⇒ S (min n m)]].
+
+let rec max n m on n ≝ match n - m with [ O => m | _ => n].
+let rec min n m on n ≝ match n - m with [ O => n | _ => m].
+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).
+
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 ≝
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 ⇒ min (v n) 1
+ | 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).
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 ≝
+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.
+
+let rec negate (F: Formula) on F ≝
match F with
[ 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)
].
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.
+
let rec dualize (F : Formula) on F : Formula ≝
match F with
[ FBot ⇒ FTop
| FNot F ⇒ FNot (dualize 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; rewrite > H; reflexivity; qed.
-
-lemma min_le: ∀m,n. n ≤ m → min n m = n.
-intro; elim m; [cases n in H; intros; [reflexivity] cases (not_le_Sn_O ? H)]
-cases n1 in H1;
-
-
-lemma xxx : ∀a,b,c.min (min a b) c = min a (min b c).
-intros; elim a; [reflexivity] elim b in c H 0; [intros;reflexivity] intros;
-lapply depth=0 (H O);
-
-
-simplify in ⊢ (? ? ? (? ? %));
-simplify in ⊢ (? ? (? % ?) ?);
-simplify in H:(? ? (? % ?) ?);
-simplify in H:(? ? ? (? ? %));
-simplify in ⊢ (? ? ? %);reflexivity]
-simplify in ⊢ (? ? (? % ?) ?);reflexivity]
+definition invert ≝
+ λv:ℕ -> ℕ. λx. if eqb (min (v x) 1) 0 then 1 else 0.
+
+(*DOCBEGIN
+
+Il linguaggio di dimostrazione di Matita
+========================================
+
+Per dimostrare questo teorema in modo agevole è necessario utilizzare il
+seguente comando:
+
+* `symmetry`
+
+ Quando la conclusuine è `a = b` permette di cambiarla in `b = a`.
+
+DOCEND*)
+theorem 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 .
+ the thesis becomes ([[ negate FBot ]]_v=[[ FBot ]]_(invert v)).
+ done.
+ case FTop .
+ the thesis becomes ([[ negate FTop ]]_v=[[ FTop ]]_(invert v)).
+ 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 (min (v n) 1 = 0 ∨ min (v n) 1 = 1) (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.
+ 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.
+ 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 H1.
+ done.
+ case FOr.
+ 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.
+ done.
+ case FImpl.
+ 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.
+ done.
+ case FNot.
+ 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.
+ done.
+qed.
+
+(*
+lemma negate_fun : ∀F,G. F ≡ G → negate F ≡ negate G.
+intros; intro v; rewrite > (negate_invert ? v);rewrite > (negate_invert ? v);
+apply H;
+qed.
*)
-lemma min_minus_n_m : ∀n,m. min (n-m) n = n-m.
-intros; apply (nat_elim2 ???? n m); intros; [reflexivity]
-[ rewrite < minus_n_O; apply min_n_n]
-
-
- cases m; [ rewrite < minus_n_O; apply min_n_n]
-
-lapply (H (S w)) as K; rewrite < minus_n_O in K; apply eq_f; assumption;]
-rewrite < (H n2) in ⊢ (? ? ? %);
-
-simplify in ⊢ (? ? (? % ?) ?);simplify; rewrite > (H n1);
-
-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 ⊢ (? ? ? (? ? %));
-
-
+theorem negate_fun:
+ ∀F:Formula.∀G:Formula.F ≡ G→negate F ≡ negate G.
+ 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).
+ assume v:(ℕ→ℕ).
+ conclude
+ [[ negate F ]]_v
+ = [[ F ]]_(invert v) by negate_invert.
+ = [[ G ]]_(invert v) by H.
+ = [[ negate G ]]_v by negate_invert.
+ done.
+qed.
+(*
+lemma not_dualize_eq_negate : ∀F. FNot (dualize F) ≡ negate F.
+intros; intro; elim F; intros; try reflexivity;
+[1,2: simplify in ⊢ (? ? ? %); rewrite <(H); rewrite <(H1);
+ [rewrite < (min_max (dualize f) (dualize f1) v); reflexivity;
+ |rewrite < (max_min (dualize f) (dualize f1) v); reflexivity;]
+|3: change in ⊢ (? ? ? %) with [[FImpl (negate f) (negate f1)]]_v;
+ change in ⊢ (? ? ? %) with (max (1 - [[negate f]]_v) [[negate f1]]_v);
+ rewrite <H1; rewrite <H;
+ rewrite > (max_min (FNot (dualize f)) ((dualize f1)) v);reflexivity;
+|4: simplify; rewrite < H; reflexivity;]
+qed.
+*)
-
+theorem not_dualize_eq_negate:
+ ∀F:Formula.negate F ≡ FNot (dualize F).
+ assume F:Formula.
+ the thesis becomes (∀v:ℕ→ℕ.[[negate F]]_v=[[FNot (dualize F)]]_v).
+ assume v:(ℕ→ℕ).
+ we proceed by induction on F to prove ([[negate F]]_v=[[FNot (dualize F)]]_v).
+ case FBot .
+ the thesis becomes ([[ negate FBot ]]_v=[[ FNot (dualize FBot) ]]_v).
+ done.
+ case FTop .
+ the thesis becomes ([[ negate FTop ]]_v=[[ FNot (dualize FTop) ]]_v).
+ done.
+ case FAtom.
+ assume n : ℕ.
+ the thesis becomes ([[ negate (FAtom n) ]]_v=[[ FNot (dualize (FAtom n)) ]]_v).
+ 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.
+ 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.
+ done.
+ case FImpl.
+ 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.
+ done.
+ case FNot.
+ 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.
+ done.
+qed.
+
+
+(*
+lemma not_inj : ∀F,G. FNot F ≡ FNot G → F ≡ G.
+intros; intro v;lapply (H v) as K;
+change in K with (1 - [[ F ]]_v = 1 - [[ G ]]_v);
+cases (sem_bool F v);cases (sem_bool G v); rewrite > H1; rewrite > H2;
+try reflexivity; rewrite > H1 in K; rewrite > H2 in K; simplify in K;
+symmetry; assumption;
+qed.
+*)
-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
+theorem not_inj:
+ ∀F:Formula.∀G:Formula.FNot F ≡ FNot G→F ≡ G.
+ 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).
+ 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 sem_bool we proved ([[ G ]]_v=O∨[[ G ]]_v=1) (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.
+ 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.
+ done.
+qed.
+
+(*
+theorem duality: ∀F1,F2. F1 ≡ F2 → dualize F1 ≡ dualize F2.
+intros; apply not_inj; intro v; rewrite > (not_dualize_eq_negate ? v);
+rewrite > (not_dualize_eq_negate ? v); apply (negate_fun ??? v); apply H;
+qed.
+*)
-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)).
-dual f = dual (subst subst f)
- = (dual subst) (subst f)
- = not (subst f)
- =
+theorem duality:
+ ∀F1:Formula.∀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 negate_fun we proved (negate F1 ≡ negate F2) (H1).
+ by not_dualize_eq_negate, equiv_rewrite we proved (FNot (dualize F1) ≡ negate F2) (H2).
+ by not_dualize_eq_negate, equiv_rewrite we proved (FNot (dualize F1) ≡ FNot (dualize F2)) (H3).
+ by not_inj we proved (dualize F1 ≡ dualize F2) (H4).
+ done.
+qed.
\ No newline at end of file