2014-11-28 15:20:52 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2014-08-12 00:35:25 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
Module: logic.examples.instances_test
|
|
|
|
|
Author: Jeremy Avigad
|
|
|
|
|
|
|
|
|
|
Illustrates substitution and congruence with iff.
|
|
|
|
|
-/
|
2014-08-12 00:35:25 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
import ..instances
|
2014-09-03 23:00:38 +00:00
|
|
|
|
open relation
|
2014-11-28 15:20:52 +00:00
|
|
|
|
open relation.general_subst
|
2014-09-03 23:00:38 +00:00
|
|
|
|
open relation.iff_ops
|
2014-10-02 00:51:17 +00:00
|
|
|
|
open eq.ops
|
2014-08-12 00:35:25 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (a b : Prop) (H : a ↔ b) (H1 : a) : b := mp H H1
|
2014-08-12 00:35:25 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (a b c d e : Prop) (H1 : a ↔ b) (H2 : a ∨ c → ¬(d → a)) : b ∨ c → ¬(d → b) :=
|
2014-08-20 02:32:44 +00:00
|
|
|
|
subst iff H1 H2
|
2014-08-12 00:35:25 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (a b c d e : Prop) (H1 : a ↔ b) (H2 : a ∨ c → ¬(d → a)) : b ∨ c → ¬(d → b) :=
|
2014-08-12 00:35:25 +00:00
|
|
|
|
H1 ▸ H2
|
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (a b c d e : Prop) (H1 : a ↔ b) : (a ∨ c → ¬(d → a)) ↔ (b ∨ c → ¬(d → b)) :=
|
|
|
|
|
is_congruence.congr iff (λa, (a ∨ c → ¬(d → a))) H1
|
2014-08-12 00:35:25 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (T : Type) (a b c d : T) (H1 : a = b) (H2 : c = b) (H3 : c = d) : a = d :=
|
2014-08-12 00:35:25 +00:00
|
|
|
|
H1 ⬝ H2⁻¹ ⬝ H3
|
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (a b c d : Prop) (H1 : a ↔ b) (H2 : c ↔ b) (H3 : c ↔ d) : a ↔ d :=
|
2014-08-12 00:35:25 +00:00
|
|
|
|
H1 ⬝ (H2⁻¹ ⬝ H3)
|
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (T : Type) (a b c d : T) (H1 : a = b) (H2 : c = b) (H3 : c = d) : a = d :=
|
2014-09-05 01:41:06 +00:00
|
|
|
|
H1 ⬝ H2⁻¹ ⬝ H3
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-11-28 15:20:52 +00:00
|
|
|
|
example (a b c d : Prop) (H1 : a ↔ b) (H2 : c ↔ b) (H3 : c ↔ d) : a ↔ d :=
|
2014-09-05 01:41:06 +00:00
|
|
|
|
H1 ⬝ H2⁻¹ ⬝ H3
|