]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/common/option_lemmas.ma
freescale porting
[helm.git] / helm / software / matita / contribs / ng_assembly / common / option_lemmas.ma
index 144733fe967c278adc678349a501f533861abc68..7e73870bf5d5fa54061829e32f5a95cce5ba9236 100644 (file)
@@ -54,10 +54,10 @@ nlemma option_destruct_none_some : ∀T.∀x:T.None T = Some T x → False.
 nqed.
 
 nlemma symmetric_eqoption :
-∀T:Type.∀op1,op2:option T.∀f:T → T → bool.
+∀T:Type.∀f:T → T → bool.∀op1,op2:option T.
  (symmetricT T bool f) →
- (eq_option T op1 op2 f = eq_option T op2 op1 f).
- #T; #op1; #op2; #f; #H;
+ (eq_option T f op1 op2 = eq_option T f op2 op1).
+ #T; #f; #op1; #op2; #H;
  nelim op1;
  nelim op2;
  nnormalize;
@@ -70,10 +70,10 @@ nlemma symmetric_eqoption :
 nqed.
 
 nlemma eq_to_eqoption :
-∀T.∀op1,op2:option T.∀f:T → T → bool.
+∀T.∀f:T → T → bool.∀op1,op2:option T.
  (∀x1,x2:T.x1 = x2 → f x1 x2 = true) →
- (op1 = op2 → eq_option T op1 op2 f = true).
- #T; #op1; #op2; #f; #H;
+ (op1 = op2 → eq_option T f op1 op2 = true).
+ #T; #f; #op1; #op2; #H;
  nelim op1;
  nelim op2;
  nnormalize;
@@ -91,10 +91,10 @@ nlemma eq_to_eqoption :
 nqed.
 
 nlemma eqoption_to_eq :
-∀T.∀op1,op2:option T.∀f:T → T → bool.
+∀T.∀f:T → T → bool.∀op1,op2:option T.
  (∀x1,x2:T.f x1 x2 = true → x1 = x2) →
- (eq_option T op1 op2 f = true → op1 = op2).
- #T; #op1; #op2; #f; #H;
+ (eq_option T f op1 op2 = true → op1 = op2).
+ #T; #f; #op1; #op2; #H;
  nelim op1;
  nelim op2;
  nnormalize;
@@ -132,17 +132,17 @@ nlemma decidable_option : ∀T.∀H:(Πx,y:T.decidable (x = y)).∀x,y:option T.
 nqed.
 
 nlemma neq_to_neqoption :
-∀T.∀op1,op2:option T.∀f:T → T → bool.
+∀T.∀f:T → T → bool.∀op1,op2:option T.
  (∀x1,x2:T.x1 ≠ x2 → f x1 x2 = false) →
- (op1 ≠ op2 → eq_option T op1 op2 f = false).
- #T; #op1; nelim op1;
+ (op1 ≠ op2 → eq_option T f op1 op2 = false).
+ #T; #f; #op1; nelim op1;
  ##[ ##1: #op2; ncases op2;
-          ##[ ##1: nnormalize; #f; #H; #H1; nelim (H1 (refl_eq …))
-          ##| ##2: #yy; #f; #H; nnormalize; #H1; napply refl_eq
+          ##[ ##1: nnormalize; #H; #H1; nelim (H1 (refl_eq …))
+          ##| ##2: #yy; #H; nnormalize; #H1; napply refl_eq
           ##]
  ##| ##2: #xx; #op2; ncases op2;
-          ##[ ##1: #f; #H; nnormalize; #H1; napply refl_eq
-          ##| ##2: #yy; #f; #H; nnormalize; #H1; napply (H xx yy …);
+          ##[ ##1: #H; nnormalize; #H1; napply refl_eq
+          ##| ##2: #yy; #H; nnormalize; #H1; napply (H xx yy …);
                    nnormalize; #H2; nrewrite > H2 in H1:(%); #H1;
                    napply (H1 (refl_eq …))
           ##]
@@ -150,22 +150,22 @@ nlemma neq_to_neqoption :
 nqed.
 
 nlemma neqoption_to_neq :
-∀T.∀op1,op2:option T.∀f:T → T → bool.
+∀T.∀f:T → T → bool.∀op1,op2:option T.
  (∀x1,x2:T.f x1 x2 = false → x1 ≠ x2) →
- (eq_option T op1 op2 f = false → op1 ≠ op2).
- #T; #op1; nelim op1;
+ (eq_option T f op1 op2 = false → op1 ≠ op2).
+ #T; #f; #op1; nelim op1;
  ##[ ##1: #op2; ncases op2;
-          ##[ ##1: nnormalize; #f; #H; #H1;
+          ##[ ##1: nnormalize; #H; #H1;
                    ndestruct (*napply (bool_destruct … H1)*)
-          ##| ##2: #yy; #f; #H; nnormalize; #H1; #H2;
+          ##| ##2: #yy; #H; nnormalize; #H1; #H2;
                    (* !!! ndestruct: assert false *)
                    napply (option_destruct_none_some T … H2)
           ##]
  ##| ##2: #xx; #op2; ncases op2;
-          ##[ ##1: nnormalize; #f; #H; #H1; #H2;
+          ##[ ##1: nnormalize; #H; #H1; #H2;
                    (* !!! ndestruct: assert false *)
                    napply (option_destruct_some_none T … H2)
-          ##| ##2: #yy; #f; #H; nnormalize; #H1; #H2; napply (H xx yy H1 ?);
+          ##| ##2: #yy; #H; nnormalize; #H1; #H2; napply (H xx yy H1 ?);
                    napply (option_destruct_some_some T … H2)
           ##]
  ##]