feat(frontends/lean): use lowercase commands, replace 'endscope' and 'endnamespace' with 'end'
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
6569b07b7c
commit
4ba097a141
261 changed files with 2020 additions and 2011 deletions
|
@ -1,9 +1,9 @@
|
|||
Import macros.
|
||||
import macros.
|
||||
|
||||
-- In this example, we prove two simple theorems about even/odd numbers.
|
||||
-- First, we define the predicates even and odd.
|
||||
Definition even (a : Nat) := ∃ b, a = 2*b.
|
||||
Definition odd (a : Nat) := ∃ b, a = 2*b + 1.
|
||||
definition even (a : Nat) := ∃ b, a = 2*b.
|
||||
definition odd (a : Nat) := ∃ b, a = 2*b + 1.
|
||||
|
||||
-- Prove that the sum of two even numbers is even.
|
||||
--
|
||||
|
@ -28,7 +28,7 @@ Definition odd (a : Nat) := ∃ b, a = 2*b + 1.
|
|||
-- We use a calculational proof 'calc' expression to derive
|
||||
-- the witness w1+w2 for the fact that a+b is also even.
|
||||
-- That is, we provide a derivation showing that a+b = 2*(w1 + w2)
|
||||
Theorem EvenPlusEven {a b : Nat} (H1 : even a) (H2 : even b) : even (a + b)
|
||||
theorem EvenPlusEven {a b : Nat} (H1 : even a) (H2 : even b) : even (a + b)
|
||||
:= Obtain (w1 : Nat) (Hw1 : a = 2*w1), from H1,
|
||||
Obtain (w2 : Nat) (Hw2 : b = 2*w2), from H2,
|
||||
ExistsIntro (w1 + w2)
|
||||
|
@ -41,7 +41,7 @@ Theorem EvenPlusEven {a b : Nat} (H1 : even a) (H2 : even b) : even (a + b)
|
|||
--
|
||||
-- Refl is the reflexivity proof. (Refl a) is a proof that two
|
||||
-- definitionally equal terms are indeed equal.
|
||||
-- "Definitionally equal" means that they have the same normal form.
|
||||
-- "definitionally equal" means that they have the same normal form.
|
||||
-- We can also view it as "Proof by computation".
|
||||
-- The normal form of (1+1), and 2*1 is 2.
|
||||
--
|
||||
|
@ -49,7 +49,7 @@ Theorem EvenPlusEven {a b : Nat} (H1 : even a) (H2 : even b) : even (a + b)
|
|||
-- The gotcha is that '2*w + 1 + 1' is actually '(2*w + 1) + 1' since +
|
||||
-- is left associative. Moreover, Lean normalizer does not use
|
||||
-- any theorems such as + associativity.
|
||||
Theorem OddPlusOne {a : Nat} (H : odd a) : even (a + 1)
|
||||
theorem OddPlusOne {a : Nat} (H : odd a) : even (a + 1)
|
||||
:= Obtain (w : Nat) (Hw : a = 2*w + 1), from H,
|
||||
ExistsIntro (w + 1)
|
||||
(calc a + 1 = 2*w + 1 + 1 : { Hw }
|
||||
|
@ -59,13 +59,13 @@ Theorem OddPlusOne {a : Nat} (H : odd a) : even (a + 1)
|
|||
|
||||
-- The following command displays the proof object produced by Lean after
|
||||
-- expanding macros, and infering implicit/missing arguments.
|
||||
print Environment 2.
|
||||
print environment 2.
|
||||
|
||||
-- By default, Lean does not display implicit arguments.
|
||||
-- The following command will force it to display them.
|
||||
SetOption pp::implicit true.
|
||||
setoption pp::implicit true.
|
||||
|
||||
print Environment 2.
|
||||
print environment 2.
|
||||
|
||||
-- As an exercise, prove that the sum of two odd numbers is even,
|
||||
-- and other similar theorems.
|
||||
|
|
|
@ -1,35 +1,35 @@
|
|||
Variable N : Type
|
||||
Variable h : N -> N -> N
|
||||
variable N : Type
|
||||
variable h : N -> N -> N
|
||||
|
||||
-- Specialize congruence theorem for h-applications
|
||||
Theorem CongrH {a1 a2 b1 b2 : N} (H1 : a1 = b1) (H2 : a2 = b2) : (h a1 a2) = (h b1 b2) :=
|
||||
theorem CongrH {a1 a2 b1 b2 : N} (H1 : a1 = b1) (H2 : a2 = b2) : (h a1 a2) = (h b1 b2) :=
|
||||
Congr (Congr (Refl h) H1) H2
|
||||
|
||||
-- Declare some variables
|
||||
Variable a : N
|
||||
Variable b : N
|
||||
Variable c : N
|
||||
Variable d : N
|
||||
Variable e : N
|
||||
variable a : N
|
||||
variable b : N
|
||||
variable c : N
|
||||
variable d : N
|
||||
variable e : N
|
||||
|
||||
-- Add axioms stating facts about these variables
|
||||
Axiom H1 : (a = b ∧ b = c) ∨ (d = c ∧ a = d)
|
||||
Axiom H2 : b = e
|
||||
axiom H1 : (a = b ∧ b = c) ∨ (d = c ∧ a = d)
|
||||
axiom H2 : b = e
|
||||
|
||||
-- Proof that (h a b) = (h c e)
|
||||
Theorem T1 : (h a b) = (h c e) :=
|
||||
theorem T1 : (h a b) = (h c e) :=
|
||||
DisjCases H1
|
||||
(λ C1, CongrH (Trans (Conjunct1 C1) (Conjunct2 C1)) H2)
|
||||
(λ C2, CongrH (Trans (Conjunct2 C2) (Conjunct1 C2)) H2)
|
||||
|
||||
-- We can use theorem T1 to prove other theorems
|
||||
Theorem T2 : (h a (h a b)) = (h a (h c e)) :=
|
||||
theorem T2 : (h a (h a b)) = (h a (h c e)) :=
|
||||
CongrH (Refl a) T1
|
||||
|
||||
-- Display the last two objects (i.e., theorems) added to the environment
|
||||
print Environment 2
|
||||
print environment 2
|
||||
|
||||
-- print implicit arguments
|
||||
SetOption lean::pp::implicit true
|
||||
SetOption pp::width 150
|
||||
print Environment 2
|
||||
setoption lean::pp::implicit true
|
||||
setoption pp::width 150
|
||||
print environment 2
|
||||
|
|
|
@ -1,19 +1,19 @@
|
|||
Import macros.
|
||||
import macros.
|
||||
|
||||
Theorem simple (p q r : Bool) : (p ⇒ q) ∧ (q ⇒ r) ⇒ p ⇒ r
|
||||
theorem simple (p q r : Bool) : (p ⇒ q) ∧ (q ⇒ r) ⇒ p ⇒ r
|
||||
:= Assume H_pq_qr H_p,
|
||||
let P_pq := Conjunct1 H_pq_qr,
|
||||
P_qr := Conjunct2 H_pq_qr,
|
||||
P_q := MP P_pq H_p
|
||||
in MP P_qr P_q.
|
||||
|
||||
SetOption pp::implicit true.
|
||||
print Environment 1.
|
||||
setoption pp::implicit true.
|
||||
print environment 1.
|
||||
|
||||
Theorem simple2 (a b c : Bool) : (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c
|
||||
theorem simple2 (a b c : Bool) : (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c
|
||||
:= Assume H_abc H_ab H_a,
|
||||
let P_b := (MP H_ab H_a),
|
||||
P_bc := (MP H_abc H_a)
|
||||
in MP P_bc P_b.
|
||||
|
||||
print Environment 1.
|
||||
print environment 1.
|
||||
|
|
|
@ -1,9 +1,9 @@
|
|||
Import macros.
|
||||
import macros.
|
||||
|
||||
Theorem and_comm (a b : Bool) : (a ∧ b) ⇒ (b ∧ a)
|
||||
theorem and_comm (a b : Bool) : (a ∧ b) ⇒ (b ∧ a)
|
||||
:= Assume H_ab, Conj (Conjunct2 H_ab) (Conjunct1 H_ab).
|
||||
|
||||
Theorem or_comm (a b : Bool) : (a ∨ b) ⇒ (b ∨ a)
|
||||
theorem or_comm (a b : Bool) : (a ∨ b) ⇒ (b ∨ a)
|
||||
:= Assume H_ab,
|
||||
DisjCases H_ab
|
||||
(λ H_a, Disj2 b H_a)
|
||||
|
@ -20,9 +20,9 @@ Theorem or_comm (a b : Bool) : (a ∨ b) ⇒ (b ∨ a)
|
|||
-- (MT H H_na) : ¬(a ⇒ b)
|
||||
-- produces
|
||||
-- a
|
||||
Theorem pierce (a b : Bool) : ((a ⇒ b) ⇒ a) ⇒ a
|
||||
theorem pierce (a b : Bool) : ((a ⇒ b) ⇒ a) ⇒ a
|
||||
:= Assume H, DisjCases (EM a)
|
||||
(λ H_a, H_a)
|
||||
(λ H_na, NotImp1 (MT H H_na)).
|
||||
|
||||
print Environment 3.
|
||||
print environment 3.
|
|
@ -1,8 +1,8 @@
|
|||
Import Int
|
||||
Import tactic
|
||||
Definition a : Nat := 10
|
||||
import Int
|
||||
import tactic
|
||||
definition a : Nat := 10
|
||||
-- Trivial indicates a "proof by evaluation"
|
||||
Theorem T1 : a > 0 := Trivial
|
||||
Theorem T2 : a - 5 > 3 := Trivial
|
||||
theorem T1 : a > 0 := Trivial
|
||||
theorem T2 : a - 5 > 3 := Trivial
|
||||
-- The next one is commented because it fails
|
||||
-- Theorem T3 : a > 11 := Trivial
|
||||
-- theorem T3 : a > 11 := Trivial
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
Push
|
||||
Theorem ReflIf (A : Type)
|
||||
scope
|
||||
theorem ReflIf (A : Type)
|
||||
(R : A → A → Bool)
|
||||
(Symmetry : Π x y, R x y → R y x)
|
||||
(Transitivity : Π x y z, R x y → R y z → R x z)
|
||||
|
@ -10,11 +10,11 @@ Push
|
|||
(fun (w : A) (H : R x w),
|
||||
let L1 : R w x := Symmetry x w H
|
||||
in Transitivity x w x H L1)
|
||||
Pop
|
||||
pop::scope
|
||||
|
||||
Push
|
||||
scope
|
||||
-- Same example but using ∀ instead of Π and ⇒ instead of →
|
||||
Theorem ReflIf (A : Type)
|
||||
theorem ReflIf (A : Type)
|
||||
(R : A → A → Bool)
|
||||
(Symmetry : ∀ x y, R x y ⇒ R y x)
|
||||
(Transitivity : ∀ x y z, R x y ⇒ R y z ⇒ R x z)
|
||||
|
@ -29,10 +29,10 @@ Push
|
|||
|
||||
-- We can make the previous example less verbose by using custom notation
|
||||
|
||||
Infixl 50 ! : ForallElim
|
||||
Infixl 30 << : MP
|
||||
infixl 50 ! : ForallElim
|
||||
infixl 30 << : MP
|
||||
|
||||
Theorem ReflIf2 (A : Type)
|
||||
theorem ReflIf2 (A : Type)
|
||||
(R : A → A → Bool)
|
||||
(Symmetry : ∀ x y, R x y ⇒ R y x)
|
||||
(Transitivity : ∀ x y z, R x y ⇒ R y z ⇒ R x z)
|
||||
|
@ -45,29 +45,29 @@ Push
|
|||
let L1 : R w x := Symmetry ! x ! w << H
|
||||
in Transitivity ! x ! w ! x << H << L1))
|
||||
|
||||
print Environment 1
|
||||
Pop
|
||||
print environment 1
|
||||
pop::scope
|
||||
|
||||
Scope
|
||||
scope
|
||||
-- Same example again.
|
||||
Variable A : Type
|
||||
Variable R : A → A → Bool
|
||||
Axiom Symmetry {x y : A} : R x y → R y x
|
||||
Axiom Transitivity {x y z : A} : R x y → R y z → R x z
|
||||
Axiom Linked (x : A) : ∃ y, R x y
|
||||
variable A : Type
|
||||
variable R : A → A → Bool
|
||||
axiom Symmetry {x y : A} : R x y → R y x
|
||||
axiom Transitivity {x y z : A} : R x y → R y z → R x z
|
||||
axiom Linked (x : A) : ∃ y, R x y
|
||||
|
||||
Theorem ReflIf (x : A) : R x x :=
|
||||
theorem ReflIf (x : A) : R x x :=
|
||||
ExistsElim (Linked x) (fun (w : A) (H : R x w),
|
||||
let L1 : R w x := Symmetry H
|
||||
in Transitivity H L1)
|
||||
|
||||
-- Even more compact proof of the same theorem
|
||||
Theorem ReflIf2 (x : A) : R x x :=
|
||||
theorem ReflIf2 (x : A) : R x x :=
|
||||
ExistsElim (Linked x) (fun w H, Transitivity H (Symmetry H))
|
||||
|
||||
-- The command EndScope exports both theorem to the main scope
|
||||
-- The command Endscope exports both theorem to the main scope
|
||||
-- The variables and axioms in the scope become parameters to both theorems.
|
||||
EndScope
|
||||
end
|
||||
|
||||
-- Display the last two theorems
|
||||
print Environment 2
|
||||
print environment 2
|
|
@ -1,14 +1,14 @@
|
|||
Import macros
|
||||
import macros
|
||||
|
||||
Definition Set (A : Type) : Type := A → Bool
|
||||
definition Set (A : Type) : Type := A → Bool
|
||||
|
||||
Definition element {A : Type} (x : A) (s : Set A) := s x
|
||||
Infix 60 ∈ : element
|
||||
definition element {A : Type} (x : A) (s : Set A) := s x
|
||||
infix 60 ∈ : element
|
||||
|
||||
Definition subset {A : Type} (s1 : Set A) (s2 : Set A) := ∀ x, x ∈ s1 ⇒ x ∈ s2
|
||||
Infix 50 ⊆ : subset
|
||||
definition subset {A : Type} (s1 : Set A) (s2 : Set A) := ∀ x, x ∈ s1 ⇒ x ∈ s2
|
||||
infix 50 ⊆ : subset
|
||||
|
||||
Theorem SubsetTrans (A : Type) : ∀ s1 s2 s3 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s3 ⇒ s1 ⊆ s3 :=
|
||||
theorem SubsetTrans (A : Type) : ∀ s1 s2 s3 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s3 ⇒ s1 ⊆ s3 :=
|
||||
take s1 s2 s3, Assume (H1 : s1 ⊆ s2) (H2 : s2 ⊆ s3),
|
||||
have s1 ⊆ s3 :
|
||||
take x, Assume Hin : x ∈ s1,
|
||||
|
@ -16,11 +16,11 @@ Theorem SubsetTrans (A : Type) : ∀ s1 s2 s3 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s3
|
|||
let L1 : x ∈ s2 := MP (Instantiate H1 x) Hin
|
||||
in MP (Instantiate H2 x) L1
|
||||
|
||||
Theorem SubsetExt (A : Type) : ∀ s1 s2 : Set A, (∀ x, x ∈ s1 = x ∈ s2) ⇒ s1 = s2 :=
|
||||
theorem SubsetExt (A : Type) : ∀ s1 s2 : Set A, (∀ x, x ∈ s1 = x ∈ s2) ⇒ s1 = s2 :=
|
||||
take s1 s2, Assume (H : ∀ x, x ∈ s1 = x ∈ s2),
|
||||
Abst (fun x, Instantiate H x)
|
||||
|
||||
Theorem SubsetAntiSymm (A : Type) : ∀ s1 s2 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s1 ⇒ s1 = s2 :=
|
||||
theorem SubsetAntiSymm (A : Type) : ∀ s1 s2 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s1 ⇒ s1 = s2 :=
|
||||
take s1 s2, Assume (H1 : s1 ⊆ s2) (H2 : s2 ⊆ s1),
|
||||
have s1 = s2 :
|
||||
MP (have (∀ x, x ∈ s1 = x ∈ s2) ⇒ s1 = s2 :
|
||||
|
@ -32,7 +32,7 @@ Theorem SubsetAntiSymm (A : Type) : ∀ s1 s2 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s1
|
|||
in ImpAntisym L1 L2)
|
||||
|
||||
-- Compact (but less readable) version of the previous theorem
|
||||
Theorem SubsetAntiSymm2 (A : Type) : ∀ s1 s2 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s1 ⇒ s1 = s2 :=
|
||||
theorem SubsetAntiSymm2 (A : Type) : ∀ s1 s2 : Set A, s1 ⊆ s2 ⇒ s2 ⊆ s1 ⇒ s1 = s2 :=
|
||||
take s1 s2, Assume H1 H2,
|
||||
MP (Instantiate (SubsetExt A) s1 s2)
|
||||
(take x, ImpAntisym (Instantiate H1 x) (Instantiate H2 x))
|
||||
|
|
|
@ -2,7 +2,7 @@
|
|||
-- "holes" that must be filled using user-defined tactics.
|
||||
|
||||
(*
|
||||
-- Import useful macros for creating tactics
|
||||
-- import useful macros for creating tactics
|
||||
import("tactic.lua")
|
||||
|
||||
-- Define a simple tactic using Lua
|
||||
|
@ -18,19 +18,19 @@ conj = conj_tac()
|
|||
--
|
||||
-- The (have [expr] by [tactic]) expression is also creating a "hole" and associating a "hint" to it.
|
||||
-- The expression [expr] after the shows is fixing the type of the "hole"
|
||||
Theorem T1 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
theorem T1 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
fun assumption : A /\ B,
|
||||
let lemma1 : A := (by auto),
|
||||
lemma2 : B := (by auto)
|
||||
in (have B /\ A by auto)
|
||||
|
||||
print Environment 1. -- print proof for the previous theorem
|
||||
print environment 1. -- print proof for the previous theorem
|
||||
|
||||
-- When hints are not provided, the user must fill the (remaining) holes using tactic command sequences.
|
||||
-- Each hole must be filled with a tactic command sequence that terminates with the command 'done' and
|
||||
-- successfully produces a proof term for filling the hole. Here is the same example without hints
|
||||
-- This style is more convenient for interactive proofs
|
||||
Theorem T2 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
theorem T2 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
fun assumption : A /\ B,
|
||||
let lemma1 : A := _, -- first hole
|
||||
lemma2 : B := _ -- second hole
|
||||
|
@ -40,7 +40,7 @@ Theorem T2 (A B : Bool) : A /\ B -> B /\ A :=
|
|||
auto. done. -- tactic command sequence for the third hole
|
||||
|
||||
-- In the following example, instead of using the "auto" tactic, we apply a sequence of even simpler tactics.
|
||||
Theorem T3 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
theorem T3 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
fun assumption : A /\ B,
|
||||
let lemma1 : A := _, -- first hole
|
||||
lemma2 : B := _ -- second hole
|
||||
|
@ -50,7 +50,7 @@ Theorem T3 (A B : Bool) : A /\ B -> B /\ A :=
|
|||
conj. exact. done. -- tactic command sequence for the third hole
|
||||
|
||||
-- We can also mix the two styles (hints and command sequences)
|
||||
Theorem T4 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
theorem T4 (A B : Bool) : A /\ B -> B /\ A :=
|
||||
fun assumption : A /\ B,
|
||||
let lemma1 : A := _, -- first hole
|
||||
lemma2 : B := _ -- second hole
|
||||
|
|
|
@ -64,9 +64,9 @@
|
|||
-- Now, the tactic conj_in_lua can be used when proving theorems in Lean.
|
||||
*)
|
||||
|
||||
Theorem T (a b : Bool) : a => b => a /\ b := _.
|
||||
theorem T (a b : Bool) : a => b => a /\ b := _.
|
||||
(* Then(Repeat(OrElse(imp_tac(), conj_in_lua)), assumption_tac()) *)
|
||||
done
|
||||
|
||||
-- print proof created using our script
|
||||
print Environment 1.
|
||||
print environment 1.
|
||||
|
|
|
@ -1,68 +1,68 @@
|
|||
Import Nat.
|
||||
import Nat
|
||||
|
||||
Variable Int : Type.
|
||||
Alias ℤ : Int.
|
||||
Builtin nat_to_int : Nat → Int.
|
||||
Coercion nat_to_int.
|
||||
variable Int : Type
|
||||
alias ℤ : Int
|
||||
builtin nat_to_int : Nat → Int
|
||||
coercion nat_to_int
|
||||
|
||||
Namespace Int.
|
||||
Builtin numeral.
|
||||
namespace Int
|
||||
builtin numeral
|
||||
|
||||
Builtin add : Int → Int → Int.
|
||||
Infixl 65 + : add.
|
||||
builtin add : Int → Int → Int
|
||||
infixl 65 + : add
|
||||
|
||||
Builtin mul : Int → Int → Int.
|
||||
Infixl 70 * : mul.
|
||||
builtin mul : Int → Int → Int
|
||||
infixl 70 * : mul
|
||||
|
||||
Builtin div : Int → Int → Int.
|
||||
Infixl 70 div : div.
|
||||
builtin div : Int → Int → Int
|
||||
infixl 70 div : div
|
||||
|
||||
Builtin le : Int → Int → Bool.
|
||||
Infix 50 <= : le.
|
||||
Infix 50 ≤ : le.
|
||||
builtin le : Int → Int → Bool
|
||||
infix 50 <= : le
|
||||
infix 50 ≤ : le
|
||||
|
||||
Definition ge (a b : Int) : Bool := b ≤ a.
|
||||
Infix 50 >= : ge.
|
||||
Infix 50 ≥ : ge.
|
||||
definition ge (a b : Int) : Bool := b ≤ a
|
||||
infix 50 >= : ge
|
||||
infix 50 ≥ : ge
|
||||
|
||||
Definition lt (a b : Int) : Bool := ¬ (a ≥ b).
|
||||
Infix 50 < : lt.
|
||||
definition lt (a b : Int) : Bool := ¬ (a ≥ b)
|
||||
infix 50 < : lt
|
||||
|
||||
Definition gt (a b : Int) : Bool := ¬ (a ≤ b).
|
||||
Infix 50 > : gt.
|
||||
definition gt (a b : Int) : Bool := ¬ (a ≤ b)
|
||||
infix 50 > : gt
|
||||
|
||||
Definition sub (a b : Int) : Int := a + -1 * b.
|
||||
Infixl 65 - : sub.
|
||||
definition sub (a b : Int) : Int := a + -1 * b
|
||||
infixl 65 - : sub
|
||||
|
||||
Definition neg (a : Int) : Int := -1 * a.
|
||||
Notation 75 - _ : neg.
|
||||
definition neg (a : Int) : Int := -1 * a
|
||||
notation 75 - _ : neg
|
||||
|
||||
Definition mod (a b : Int) : Int := a - b * (a div b).
|
||||
Infixl 70 mod : mod.
|
||||
definition mod (a b : Int) : Int := a - b * (a div b)
|
||||
infixl 70 mod : mod
|
||||
|
||||
Definition divides (a b : Int) : Bool := (b mod a) = 0.
|
||||
Infix 50 | : divides.
|
||||
definition divides (a b : Int) : Bool := (b mod a) = 0
|
||||
infix 50 | : divides
|
||||
|
||||
Definition abs (a : Int) : Int := if (0 ≤ a) a (- a).
|
||||
Notation 55 | _ | : abs.
|
||||
definition abs (a : Int) : Int := if (0 ≤ a) a (- a)
|
||||
notation 55 | _ | : abs
|
||||
|
||||
SetOpaque sub true.
|
||||
SetOpaque neg true.
|
||||
SetOpaque mod true.
|
||||
SetOpaque divides true.
|
||||
SetOpaque abs true.
|
||||
SetOpaque ge true.
|
||||
SetOpaque lt true.
|
||||
SetOpaque gt true.
|
||||
EndNamespace.
|
||||
setopaque sub true
|
||||
setopaque neg true
|
||||
setopaque mod true
|
||||
setopaque divides true
|
||||
setopaque abs true
|
||||
setopaque ge true
|
||||
setopaque lt true
|
||||
setopaque gt true
|
||||
end
|
||||
|
||||
Namespace Nat.
|
||||
Definition sub (a b : Nat) : Int := nat_to_int a - nat_to_int b.
|
||||
Infixl 65 - : sub.
|
||||
namespace Nat
|
||||
definition sub (a b : Nat) : Int := nat_to_int a - nat_to_int b
|
||||
infixl 65 - : sub
|
||||
|
||||
Definition neg (a : Nat) : Int := - (nat_to_int a).
|
||||
Notation 75 - _ : neg.
|
||||
definition neg (a : Nat) : Int := - (nat_to_int a)
|
||||
notation 75 - _ : neg
|
||||
|
||||
SetOpaque sub true.
|
||||
SetOpaque neg true.
|
||||
EndNamespace.
|
||||
setopaque sub true
|
||||
setopaque neg true
|
||||
end
|
|
@ -1,47 +1,47 @@
|
|||
Import kernel.
|
||||
Import macros.
|
||||
import kernel
|
||||
import macros
|
||||
|
||||
Variable Nat : Type.
|
||||
Alias ℕ : Nat.
|
||||
variable Nat : Type
|
||||
alias ℕ : Nat
|
||||
|
||||
Namespace Nat.
|
||||
Builtin numeral.
|
||||
namespace Nat
|
||||
builtin numeral
|
||||
|
||||
Builtin add : Nat → Nat → Nat.
|
||||
Infixl 65 + : add.
|
||||
builtin add : Nat → Nat → Nat
|
||||
infixl 65 + : add
|
||||
|
||||
Builtin mul : Nat → Nat → Nat.
|
||||
Infixl 70 * : mul.
|
||||
builtin mul : Nat → Nat → Nat
|
||||
infixl 70 * : mul
|
||||
|
||||
Builtin le : Nat → Nat → Bool.
|
||||
Infix 50 <= : le.
|
||||
Infix 50 ≤ : le.
|
||||
builtin le : Nat → Nat → Bool
|
||||
infix 50 <= : le
|
||||
infix 50 ≤ : le
|
||||
|
||||
Definition ge (a b : Nat) := b ≤ a.
|
||||
Infix 50 >= : ge.
|
||||
Infix 50 ≥ : ge.
|
||||
definition ge (a b : Nat) := b ≤ a
|
||||
infix 50 >= : ge
|
||||
infix 50 ≥ : ge
|
||||
|
||||
Definition lt (a b : Nat) := ¬ (a ≥ b).
|
||||
Infix 50 < : lt.
|
||||
definition lt (a b : Nat) := ¬ (a ≥ b)
|
||||
infix 50 < : lt
|
||||
|
||||
Definition gt (a b : Nat) := ¬ (a ≤ b).
|
||||
Infix 50 > : gt.
|
||||
definition gt (a b : Nat) := ¬ (a ≤ b)
|
||||
infix 50 > : gt
|
||||
|
||||
Definition id (a : Nat) := a.
|
||||
Notation 55 | _ | : id.
|
||||
definition id (a : Nat) := a
|
||||
notation 55 | _ | : id
|
||||
|
||||
Axiom SuccNeZero (a : Nat) : a + 1 ≠ 0.
|
||||
Axiom SuccInj {a b : Nat} (H : a + 1 = b + 1) : a = b
|
||||
Axiom PlusZero (a : Nat) : a + 0 = a.
|
||||
Axiom PlusSucc (a b : Nat) : a + (b + 1) = (a + b) + 1.
|
||||
Axiom MulZero (a : Nat) : a * 0 = 0.
|
||||
Axiom MulSucc (a b : Nat) : a * (b + 1) = a * b + a.
|
||||
Axiom LeDef (a b : Nat) : a ≤ b ⇔ ∃ c, a + c = b.
|
||||
Axiom Induction {P : Nat → Bool} (a : Nat) (H1 : P 0) (H2 : Π (n : Nat) (iH : P n), P (n + 1)) : P a.
|
||||
axiom SuccNeZero (a : Nat) : a + 1 ≠ 0
|
||||
axiom SuccInj {a b : Nat} (H : a + 1 = b + 1) : a = b
|
||||
axiom PlusZero (a : Nat) : a + 0 = a
|
||||
axiom PlusSucc (a b : Nat) : a + (b + 1) = (a + b) + 1
|
||||
axiom MulZero (a : Nat) : a * 0 = 0
|
||||
axiom MulSucc (a b : Nat) : a * (b + 1) = a * b + a
|
||||
axiom LeDef (a b : Nat) : a ≤ b ⇔ ∃ c, a + c = b
|
||||
axiom Induction {P : Nat → Bool} (a : Nat) (H1 : P 0) (H2 : Π (n : Nat) (iH : P n), P (n + 1)) : P a
|
||||
|
||||
Theorem ZeroNeOne : 0 ≠ 1 := Trivial.
|
||||
theorem ZeroNeOne : 0 ≠ 1 := Trivial
|
||||
|
||||
Theorem NeZeroPred' (a : Nat) : a ≠ 0 ⇒ ∃ b, b + 1 = a
|
||||
theorem NeZeroPred' (a : Nat) : a ≠ 0 ⇒ ∃ b, b + 1 = a
|
||||
:= Induction a
|
||||
(Assume H : 0 ≠ 0, FalseElim (∃ b, b + 1 = 0) H)
|
||||
(λ (n : Nat) (iH : n ≠ 0 ⇒ ∃ b, b + 1 = n),
|
||||
|
@ -50,99 +50,99 @@ Theorem NeZeroPred' (a : Nat) : a ≠ 0 ⇒ ∃ b, b + 1 = a
|
|||
(λ Heq0 : n = 0, ExistsIntro 0 (calc 0 + 1 = n + 1 : { Symm Heq0 }))
|
||||
(λ Hne0 : n ≠ 0,
|
||||
Obtain (w : Nat) (Hw : w + 1 = n), from (MP iH Hne0),
|
||||
ExistsIntro (w + 1) (calc w + 1 + 1 = n + 1 : { Hw }))).
|
||||
ExistsIntro (w + 1) (calc w + 1 + 1 = n + 1 : { Hw })))
|
||||
|
||||
Theorem NeZeroPred {a : Nat} (H : a ≠ 0) : ∃ b, b + 1 = a
|
||||
:= MP (NeZeroPred' a) H.
|
||||
theorem NeZeroPred {a : Nat} (H : a ≠ 0) : ∃ b, b + 1 = a
|
||||
:= MP (NeZeroPred' a) H
|
||||
|
||||
Theorem Destruct {B : Bool} {a : Nat} (H1: a = 0 → B) (H2 : Π n, a = n + 1 → B) : B
|
||||
theorem Destruct {B : Bool} {a : Nat} (H1: a = 0 → B) (H2 : Π n, a = n + 1 → B) : B
|
||||
:= DisjCases (EM (a = 0))
|
||||
(λ Heq0 : a = 0, H1 Heq0)
|
||||
(λ Hne0 : a ≠ 0, Obtain (w : Nat) (Hw : w + 1 = a), from (NeZeroPred Hne0),
|
||||
H2 w (Symm Hw)).
|
||||
H2 w (Symm Hw))
|
||||
|
||||
Theorem ZeroPlus (a : Nat) : 0 + a = a
|
||||
theorem ZeroPlus (a : Nat) : 0 + a = a
|
||||
:= Induction a
|
||||
(have 0 + 0 = 0 : Trivial)
|
||||
(λ (n : Nat) (iH : 0 + n = n),
|
||||
calc 0 + (n + 1) = (0 + n) + 1 : PlusSucc 0 n
|
||||
... = n + 1 : { iH }).
|
||||
... = n + 1 : { iH })
|
||||
|
||||
Theorem SuccPlus (a b : Nat) : (a + 1) + b = (a + b) + 1
|
||||
theorem SuccPlus (a b : Nat) : (a + 1) + b = (a + b) + 1
|
||||
:= Induction b
|
||||
(calc (a + 1) + 0 = a + 1 : PlusZero (a + 1)
|
||||
... = (a + 0) + 1 : { Symm (PlusZero a) })
|
||||
... = (a + 0) + 1 : { Symm (PlusZero a) })
|
||||
(λ (n : Nat) (iH : (a + 1) + n = (a + n) + 1),
|
||||
calc (a + 1) + (n + 1) = ((a + 1) + n) + 1 : PlusSucc (a + 1) n
|
||||
... = ((a + n) + 1) + 1 : { iH }
|
||||
... = (a + (n + 1)) + 1 : { have (a + n) + 1 = a + (n + 1) : Symm (PlusSucc a n) }).
|
||||
... = ((a + n) + 1) + 1 : { iH }
|
||||
... = (a + (n + 1)) + 1 : { have (a + n) + 1 = a + (n + 1) : Symm (PlusSucc a n) })
|
||||
|
||||
Theorem PlusComm (a b : Nat) : a + b = b + a
|
||||
theorem PlusComm (a b : Nat) : a + b = b + a
|
||||
:= Induction b
|
||||
(calc a + 0 = a : PlusZero a
|
||||
... = 0 + a : Symm (ZeroPlus a))
|
||||
(λ (n : Nat) (iH : a + n = n + a),
|
||||
calc a + (n + 1) = (a + n) + 1 : PlusSucc a n
|
||||
... = (n + a) + 1 : { iH }
|
||||
... = (n + 1) + a : Symm (SuccPlus n a)).
|
||||
... = (n + a) + 1 : { iH }
|
||||
... = (n + 1) + a : Symm (SuccPlus n a))
|
||||
|
||||
Theorem PlusAssoc (a b c : Nat) : a + (b + c) = (a + b) + c
|
||||
theorem PlusAssoc (a b c : Nat) : a + (b + c) = (a + b) + c
|
||||
:= Induction a
|
||||
(calc 0 + (b + c) = b + c : ZeroPlus (b + c)
|
||||
... = (0 + b) + c : { Symm (ZeroPlus b) })
|
||||
... = (0 + b) + c : { Symm (ZeroPlus b) })
|
||||
(λ (n : Nat) (iH : n + (b + c) = (n + b) + c),
|
||||
calc (n + 1) + (b + c) = (n + (b + c)) + 1 : SuccPlus n (b + c)
|
||||
... = ((n + b) + c) + 1 : { iH }
|
||||
... = ((n + b) + 1) + c : Symm (SuccPlus (n + b) c)
|
||||
... = ((n + 1) + b) + c : { have (n + b) + 1 = (n + 1) + b : Symm (SuccPlus n b) }).
|
||||
... = ((n + 1) + b) + c : { have (n + b) + 1 = (n + 1) + b : Symm (SuccPlus n b) })
|
||||
|
||||
Theorem ZeroMul (a : Nat) : 0 * a = 0
|
||||
theorem ZeroMul (a : Nat) : 0 * a = 0
|
||||
:= Induction a
|
||||
(have 0 * 0 = 0 : Trivial)
|
||||
(λ (n : Nat) (iH : 0 * n = 0),
|
||||
calc 0 * (n + 1) = (0 * n) + 0 : MulSucc 0 n
|
||||
... = 0 + 0 : { iH }
|
||||
... = 0 : Trivial).
|
||||
... = 0 : Trivial)
|
||||
|
||||
Theorem SuccMul (a b : Nat) : (a + 1) * b = a * b + b
|
||||
theorem SuccMul (a b : Nat) : (a + 1) * b = a * b + b
|
||||
:= Induction b
|
||||
(calc (a + 1) * 0 = 0 : MulZero (a + 1)
|
||||
... = a * 0 : Symm (MulZero a)
|
||||
... = a * 0 + 0 : Symm (PlusZero (a * 0)))
|
||||
... = a * 0 : Symm (MulZero a)
|
||||
... = a * 0 + 0 : Symm (PlusZero (a * 0)))
|
||||
(λ (n : Nat) (iH : (a + 1) * n = a * n + n),
|
||||
calc (a + 1) * (n + 1) = (a + 1) * n + (a + 1) : MulSucc (a + 1) n
|
||||
... = a * n + n + (a + 1) : { iH }
|
||||
... = a * n + n + a + 1 : PlusAssoc (a * n + n) a 1
|
||||
... = a * n + (n + a) + 1 : { have a * n + n + a = a * n + (n + a) : Symm (PlusAssoc (a * n) n a) }
|
||||
... = a * n + (a + n) + 1 : { PlusComm n a }
|
||||
... = a * n + a + n + 1 : { PlusAssoc (a * n) a n }
|
||||
... = a * (n + 1) + n + 1 : { Symm (MulSucc a n) }
|
||||
... = a * (n + 1) + (n + 1) : Symm (PlusAssoc (a * (n + 1)) n 1)).
|
||||
... = a * n + n + (a + 1) : { iH }
|
||||
... = a * n + n + a + 1 : PlusAssoc (a * n + n) a 1
|
||||
... = a * n + (n + a) + 1 : { have a * n + n + a = a * n + (n + a) : Symm (PlusAssoc (a * n) n a) }
|
||||
... = a * n + (a + n) + 1 : { PlusComm n a }
|
||||
... = a * n + a + n + 1 : { PlusAssoc (a * n) a n }
|
||||
... = a * (n + 1) + n + 1 : { Symm (MulSucc a n) }
|
||||
... = a * (n + 1) + (n + 1) : Symm (PlusAssoc (a * (n + 1)) n 1))
|
||||
|
||||
Theorem OneMul (a : Nat) : 1 * a = a
|
||||
theorem OneMul (a : Nat) : 1 * a = a
|
||||
:= Induction a
|
||||
(have 1 * 0 = 0 : Trivial)
|
||||
(λ (n : Nat) (iH : 1 * n = n),
|
||||
calc 1 * (n + 1) = 1 * n + 1 : MulSucc 1 n
|
||||
... = n + 1 : { iH }).
|
||||
... = n + 1 : { iH })
|
||||
|
||||
Theorem MulOne (a : Nat) : a * 1 = a
|
||||
theorem MulOne (a : Nat) : a * 1 = a
|
||||
:= Induction a
|
||||
(have 0 * 1 = 0 : Trivial)
|
||||
(λ (n : Nat) (iH : n * 1 = n),
|
||||
calc (n + 1) * 1 = n * 1 + 1 : SuccMul n 1
|
||||
... = n + 1 : { iH }).
|
||||
... = n + 1 : { iH })
|
||||
|
||||
Theorem MulComm (a b : Nat) : a * b = b * a
|
||||
theorem MulComm (a b : Nat) : a * b = b * a
|
||||
:= Induction b
|
||||
(calc a * 0 = 0 : MulZero a
|
||||
... = 0 * a : Symm (ZeroMul a))
|
||||
(λ (n : Nat) (iH : a * n = n * a),
|
||||
calc a * (n + 1) = a * n + a : MulSucc a n
|
||||
... = n * a + a : { iH }
|
||||
... = (n + 1) * a : Symm (SuccMul n a)).
|
||||
... = (n + 1) * a : Symm (SuccMul n a))
|
||||
|
||||
Theorem Distribute (a b c : Nat) : a * (b + c) = a * b + a * c
|
||||
theorem Distribute (a b c : Nat) : a * (b + c) = a * b + a * c
|
||||
:= Induction a
|
||||
(calc 0 * (b + c) = 0 : ZeroMul (b + c)
|
||||
... = 0 + 0 : Trivial
|
||||
|
@ -150,55 +150,55 @@ Theorem Distribute (a b c : Nat) : a * (b + c) = a * b + a * c
|
|||
... = 0 * b + 0 * c : { Symm (ZeroMul c) })
|
||||
(λ (n : Nat) (iH : n * (b + c) = n * b + n * c),
|
||||
calc (n + 1) * (b + c) = n * (b + c) + (b + c) : SuccMul n (b + c)
|
||||
... = n * b + n * c + (b + c) : { iH }
|
||||
... = n * b + n * c + b + c : PlusAssoc (n * b + n * c) b c
|
||||
... = n * b + (n * c + b) + c : { Symm (PlusAssoc (n * b) (n * c) b) }
|
||||
... = n * b + (b + n * c) + c : { PlusComm (n * c) b }
|
||||
... = n * b + b + n * c + c : { PlusAssoc (n * b) b (n * c) }
|
||||
... = (n + 1) * b + n * c + c : { Symm (SuccMul n b) }
|
||||
... = (n + 1) * b + (n * c + c) : Symm (PlusAssoc ((n + 1) * b) (n * c) c)
|
||||
... = (n + 1) * b + (n + 1) * c : { Symm (SuccMul n c) }).
|
||||
... = n * b + n * c + (b + c) : { iH }
|
||||
... = n * b + n * c + b + c : PlusAssoc (n * b + n * c) b c
|
||||
... = n * b + (n * c + b) + c : { Symm (PlusAssoc (n * b) (n * c) b) }
|
||||
... = n * b + (b + n * c) + c : { PlusComm (n * c) b }
|
||||
... = n * b + b + n * c + c : { PlusAssoc (n * b) b (n * c) }
|
||||
... = (n + 1) * b + n * c + c : { Symm (SuccMul n b) }
|
||||
... = (n + 1) * b + (n * c + c) : Symm (PlusAssoc ((n + 1) * b) (n * c) c)
|
||||
... = (n + 1) * b + (n + 1) * c : { Symm (SuccMul n c) })
|
||||
|
||||
Theorem Distribute2 (a b c : Nat) : (a + b) * c = a * c + b * c
|
||||
theorem Distribute2 (a b c : Nat) : (a + b) * c = a * c + b * c
|
||||
:= calc (a + b) * c = c * (a + b) : MulComm (a + b) c
|
||||
... = c * a + c * b : Distribute c a b
|
||||
... = a * c + c * b : { MulComm c a }
|
||||
... = a * c + b * c : { MulComm c b }.
|
||||
... = a * c + b * c : { MulComm c b }
|
||||
|
||||
Theorem MulAssoc (a b c : Nat) : a * (b * c) = a * b * c
|
||||
theorem MulAssoc (a b c : Nat) : a * (b * c) = a * b * c
|
||||
:= Induction a
|
||||
(calc 0 * (b * c) = 0 : ZeroMul (b * c)
|
||||
... = 0 * c : Symm (ZeroMul c)
|
||||
... = (0 * b) * c : { Symm (ZeroMul b) })
|
||||
(λ (n : Nat) (iH : n * (b * c) = n * b * c),
|
||||
calc (n + 1) * (b * c) = n * (b * c) + (b * c) : SuccMul n (b * c)
|
||||
... = n * b * c + (b * c) : { iH }
|
||||
... = (n * b + b) * c : Symm (Distribute2 (n * b) b c)
|
||||
... = (n + 1) * b * c : { Symm (SuccMul n b) }).
|
||||
... = n * b * c + (b * c) : { iH }
|
||||
... = (n * b + b) * c : Symm (Distribute2 (n * b) b c)
|
||||
... = (n + 1) * b * c : { Symm (SuccMul n b) })
|
||||
|
||||
Theorem PlusInj' (a b c : Nat) : a + b = a + c ⇒ b = c
|
||||
theorem PlusInj' (a b c : Nat) : a + b = a + c ⇒ b = c
|
||||
:= Induction a
|
||||
(Assume H : 0 + b = 0 + c,
|
||||
calc b = 0 + b : Symm (ZeroPlus b)
|
||||
... = 0 + c : H
|
||||
... = c : ZeroPlus c)
|
||||
... = 0 + c : H
|
||||
... = c : ZeroPlus c)
|
||||
(λ (n : Nat) (iH : n + b = n + c ⇒ b = c),
|
||||
Assume H : n + 1 + b = n + 1 + c,
|
||||
let L1 : n + b + 1 = n + c + 1
|
||||
:= (calc n + b + 1 = n + (b + 1) : Symm (PlusAssoc n b 1)
|
||||
... = n + (1 + b) : { PlusComm b 1 }
|
||||
... = n + 1 + b : PlusAssoc n 1 b
|
||||
... = n + 1 + c : H
|
||||
... = n + (1 + c) : Symm (PlusAssoc n 1 c)
|
||||
... = n + (c + 1) : { PlusComm 1 c }
|
||||
... = n + c + 1 : PlusAssoc n c 1),
|
||||
... = n + (1 + b) : { PlusComm b 1 }
|
||||
... = n + 1 + b : PlusAssoc n 1 b
|
||||
... = n + 1 + c : H
|
||||
... = n + (1 + c) : Symm (PlusAssoc n 1 c)
|
||||
... = n + (c + 1) : { PlusComm 1 c }
|
||||
... = n + c + 1 : PlusAssoc n c 1),
|
||||
L2 : n + b = n + c := SuccInj L1
|
||||
in MP iH L2).
|
||||
in MP iH L2)
|
||||
|
||||
Theorem PlusInj {a b c : Nat} (H : a + b = a + c) : b = c
|
||||
:= MP (PlusInj' a b c) H.
|
||||
theorem PlusInj {a b c : Nat} (H : a + b = a + c) : b = c
|
||||
:= MP (PlusInj' a b c) H
|
||||
|
||||
Theorem PlusEq0 {a b : Nat} (H : a + b = 0) : a = 0
|
||||
theorem PlusEq0 {a b : Nat} (H : a + b = 0) : a = 0
|
||||
:= Destruct (λ H1 : a = 0, H1)
|
||||
(λ (n : Nat) (H1 : a = n + 1),
|
||||
AbsurdElim (a = 0)
|
||||
|
@ -208,45 +208,45 @@ Theorem PlusEq0 {a b : Nat} (H : a + b = 0) : a = 0
|
|||
... = n + b + 1 : PlusAssoc n b 1
|
||||
... ≠ 0 : SuccNeZero (n + b)))
|
||||
|
||||
Theorem LeIntro {a b c : Nat} (H : a + c = b) : a ≤ b
|
||||
:= EqMP (Symm (LeDef a b)) (have (∃ x, a + x = b) : ExistsIntro c H).
|
||||
theorem LeIntro {a b c : Nat} (H : a + c = b) : a ≤ b
|
||||
:= EqMP (Symm (LeDef a b)) (have (∃ x, a + x = b) : ExistsIntro c H)
|
||||
|
||||
Theorem LeElim {a b : Nat} (H : a ≤ b) : ∃ x, a + x = b
|
||||
:= EqMP (LeDef a b) H.
|
||||
theorem LeElim {a b : Nat} (H : a ≤ b) : ∃ x, a + x = b
|
||||
:= EqMP (LeDef a b) H
|
||||
|
||||
Theorem LeRefl (a : Nat) : a ≤ a := LeIntro (PlusZero a).
|
||||
theorem LeRefl (a : Nat) : a ≤ a := LeIntro (PlusZero a)
|
||||
|
||||
Theorem LeZero (a : Nat) : 0 ≤ a := LeIntro (ZeroPlus a).
|
||||
theorem LeZero (a : Nat) : 0 ≤ a := LeIntro (ZeroPlus a)
|
||||
|
||||
Theorem LeTrans {a b c : Nat} (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c
|
||||
theorem LeTrans {a b c : Nat} (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c
|
||||
:= Obtain (w1 : Nat) (Hw1 : a + w1 = b), from (LeElim H1),
|
||||
Obtain (w2 : Nat) (Hw2 : b + w2 = c), from (LeElim H2),
|
||||
LeIntro (calc a + (w1 + w2) = a + w1 + w2 : PlusAssoc a w1 w2
|
||||
... = b + w2 : { Hw1 }
|
||||
... = c : Hw2).
|
||||
... = b + w2 : { Hw1 }
|
||||
... = c : Hw2)
|
||||
|
||||
Theorem LeInj {a b : Nat} (H : a ≤ b) (c : Nat) : a + c ≤ b + c
|
||||
theorem LeInj {a b : Nat} (H : a ≤ b) (c : Nat) : a + c ≤ b + c
|
||||
:= Obtain (w : Nat) (Hw : a + w = b), from (LeElim H),
|
||||
LeIntro (calc a + c + w = a + (c + w) : Symm (PlusAssoc a c w)
|
||||
... = a + (w + c) : { PlusComm c w }
|
||||
... = a + w + c : PlusAssoc a w c
|
||||
... = b + c : { Hw }).
|
||||
... = a + (w + c) : { PlusComm c w }
|
||||
... = a + w + c : PlusAssoc a w c
|
||||
... = b + c : { Hw })
|
||||
|
||||
Theorem LeAntiSymm {a b : Nat} (H1 : a ≤ b) (H2 : b ≤ a) : a = b
|
||||
theorem LeAntiSymm {a b : Nat} (H1 : a ≤ b) (H2 : b ≤ a) : a = b
|
||||
:= Obtain (w1 : Nat) (Hw1 : a + w1 = b), from (LeElim H1),
|
||||
Obtain (w2 : Nat) (Hw2 : b + w2 = a), from (LeElim H2),
|
||||
let L1 : w1 + w2 = 0
|
||||
:= PlusInj (calc a + (w1 + w2) = a + w1 + w2 : { PlusAssoc a w1 w2 }
|
||||
... = b + w2 : { Hw1 }
|
||||
... = a : Hw2
|
||||
... = a + 0 : Symm (PlusZero a)),
|
||||
... = b + w2 : { Hw1 }
|
||||
... = a : Hw2
|
||||
... = a + 0 : Symm (PlusZero a)),
|
||||
L2 : w1 = 0 := PlusEq0 L1
|
||||
in calc a = a + 0 : Symm (PlusZero a)
|
||||
... = a + w1 : { Symm L2 }
|
||||
... = b : Hw1.
|
||||
... = a + w1 : { Symm L2 }
|
||||
... = b : Hw1
|
||||
|
||||
SetOpaque ge true.
|
||||
SetOpaque lt true.
|
||||
SetOpaque gt true.
|
||||
SetOpaque id true.
|
||||
EndNamespace.
|
||||
setopaque ge true
|
||||
setopaque lt true
|
||||
setopaque gt true
|
||||
setopaque id true
|
||||
end
|
|
@ -1,44 +1,44 @@
|
|||
Import Int.
|
||||
import Int
|
||||
|
||||
Variable Real : Type.
|
||||
Alias ℝ : Real.
|
||||
Builtin int_to_real : Int → Real.
|
||||
Coercion int_to_real.
|
||||
Definition nat_to_real (a : Nat) : Real := int_to_real (nat_to_int a).
|
||||
Coercion nat_to_real.
|
||||
variable Real : Type
|
||||
alias ℝ : Real
|
||||
builtin int_to_real : Int → Real
|
||||
coercion int_to_real
|
||||
definition nat_to_real (a : Nat) : Real := int_to_real (nat_to_int a)
|
||||
coercion nat_to_real
|
||||
|
||||
Namespace Real.
|
||||
Builtin numeral.
|
||||
namespace Real
|
||||
builtin numeral
|
||||
|
||||
Builtin add : Real → Real → Real.
|
||||
Infixl 65 + : add.
|
||||
builtin add : Real → Real → Real
|
||||
infixl 65 + : add
|
||||
|
||||
Builtin mul : Real → Real → Real.
|
||||
Infixl 70 * : mul.
|
||||
builtin mul : Real → Real → Real
|
||||
infixl 70 * : mul
|
||||
|
||||
Builtin div : Real → Real → Real.
|
||||
Infixl 70 / : div.
|
||||
builtin div : Real → Real → Real
|
||||
infixl 70 / : div
|
||||
|
||||
Builtin le : Real → Real → Bool.
|
||||
Infix 50 <= : le.
|
||||
Infix 50 ≤ : le.
|
||||
builtin le : Real → Real → Bool
|
||||
infix 50 <= : le
|
||||
infix 50 ≤ : le
|
||||
|
||||
Definition ge (a b : Real) : Bool := b ≤ a.
|
||||
Infix 50 >= : ge.
|
||||
Infix 50 ≥ : ge.
|
||||
definition ge (a b : Real) : Bool := b ≤ a
|
||||
infix 50 >= : ge
|
||||
infix 50 ≥ : ge
|
||||
|
||||
Definition lt (a b : Real) : Bool := ¬ (a ≥ b).
|
||||
Infix 50 < : lt.
|
||||
definition lt (a b : Real) : Bool := ¬ (a ≥ b)
|
||||
infix 50 < : lt
|
||||
|
||||
Definition gt (a b : Real) : Bool := ¬ (a ≤ b).
|
||||
Infix 50 > : gt.
|
||||
definition gt (a b : Real) : Bool := ¬ (a ≤ b)
|
||||
infix 50 > : gt
|
||||
|
||||
Definition sub (a b : Real) : Real := a + -1.0 * b.
|
||||
Infixl 65 - : sub.
|
||||
definition sub (a b : Real) : Real := a + -1.0 * b
|
||||
infixl 65 - : sub
|
||||
|
||||
Definition neg (a : Real) : Real := -1.0 * a.
|
||||
Notation 75 - _ : neg.
|
||||
definition neg (a : Real) : Real := -1.0 * a
|
||||
notation 75 - _ : neg
|
||||
|
||||
Definition abs (a : Real) : Real := if (0.0 ≤ a) a (- a).
|
||||
Notation 55 | _ | : abs.
|
||||
EndNamespace.
|
||||
definition abs (a : Real) : Real := if (0.0 ≤ a) a (- a)
|
||||
notation 55 | _ | : abs
|
||||
end
|
||||
|
|
|
@ -2,23 +2,23 @@
|
|||
|
||||
-- The cast operator allows us to cast an element of type A
|
||||
-- into B if we provide a proof that types A and B are equal.
|
||||
Variable cast {A B : (Type U)} : A == B → A → B.
|
||||
variable cast {A B : (Type U)} : A == B → A → B.
|
||||
|
||||
-- The CastEq axiom states that for any cast of x is equal to x.
|
||||
Axiom CastEq {A B : (Type U)} (H : A == B) (x : A) : x == cast H x.
|
||||
axiom CastEq {A B : (Type U)} (H : A == B) (x : A) : x == cast H x.
|
||||
|
||||
-- The CastApp axiom "propagates" the cast over application
|
||||
Axiom CastApp {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)}
|
||||
axiom CastApp {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)}
|
||||
(H1 : (Π x : A, B x) == (Π x : A', B' x)) (H2 : A == A')
|
||||
(f : Π x : A, B x) (x : A) :
|
||||
cast H1 f (cast H2 x) == f x.
|
||||
|
||||
-- If two (dependent) function spaces are equal, then their domains are equal.
|
||||
Axiom DomInj {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)}
|
||||
axiom DomInj {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)}
|
||||
(H : (Π x : A, B x) == (Π x : A', B' x)) :
|
||||
A == A'.
|
||||
|
||||
-- If two (dependent) function spaces are equal, then their ranges are equal.
|
||||
Axiom RanInj {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)}
|
||||
axiom RanInj {A A' : (Type U)} {B : A → (Type U)} {B' : A' → (Type U)}
|
||||
(H : (Π x : A, B x) == (Π x : A', B' x)) (a : A) :
|
||||
B a == B' (cast (DomInj H) a).
|
||||
|
|
|
@ -6,7 +6,7 @@
|
|||
-- Example: the command
|
||||
-- Find "^[cC]on"
|
||||
-- Displays all objects that start with the string "con" or "Con"
|
||||
cmd_macro("Find",
|
||||
cmd_macro("find",
|
||||
{ macro_arg.String },
|
||||
function(env, pattern)
|
||||
local opts = get_options()
|
||||
|
|
|
@ -1,396 +1,396 @@
|
|||
Import macros
|
||||
import macros
|
||||
|
||||
Universe M : 512.
|
||||
Universe U : M+512.
|
||||
universe M : 512
|
||||
universe U : M+512
|
||||
|
||||
Variable Bool : Type.
|
||||
-- The following builtin declarations can be removed as soon as Lean supports inductive datatypes and match expressions.
|
||||
Builtin true : Bool.
|
||||
Builtin false : Bool.
|
||||
Builtin if {A : (Type U)} : Bool → A → A → A.
|
||||
variable Bool : Type
|
||||
-- The following builtin declarations can be removed as soon as Lean supports inductive datatypes and match expressions
|
||||
builtin true : Bool
|
||||
builtin false : Bool
|
||||
builtin if {A : (Type U)} : Bool → A → A → A
|
||||
|
||||
Definition TypeU := (Type U)
|
||||
Definition TypeM := (Type M)
|
||||
definition TypeU := (Type U)
|
||||
definition TypeM := (Type M)
|
||||
|
||||
Definition implies (a b : Bool) : Bool
|
||||
:= if a b true.
|
||||
definition implies (a b : Bool) : Bool
|
||||
:= if a b true
|
||||
|
||||
Infixr 25 => : implies.
|
||||
Infixr 25 ⇒ : implies.
|
||||
infixr 25 => : implies
|
||||
infixr 25 ⇒ : implies
|
||||
|
||||
Definition iff (a b : Bool) : Bool
|
||||
:= a == b.
|
||||
definition iff (a b : Bool) : Bool
|
||||
:= a == b
|
||||
|
||||
Infixr 25 <=> : iff.
|
||||
Infixr 25 ⇔ : iff.
|
||||
infixr 25 <=> : iff
|
||||
infixr 25 ⇔ : iff
|
||||
|
||||
Definition not (a : Bool) : Bool
|
||||
:= if a false true.
|
||||
definition not (a : Bool) : Bool
|
||||
:= if a false true
|
||||
|
||||
Notation 40 ¬ _ : not.
|
||||
notation 40 ¬ _ : not
|
||||
|
||||
Definition or (a b : Bool) : Bool
|
||||
:= ¬ a ⇒ b.
|
||||
definition or (a b : Bool) : Bool
|
||||
:= ¬ a ⇒ b
|
||||
|
||||
Infixr 30 || : or.
|
||||
Infixr 30 \/ : or.
|
||||
Infixr 30 ∨ : or.
|
||||
infixr 30 || : or
|
||||
infixr 30 \/ : or
|
||||
infixr 30 ∨ : or
|
||||
|
||||
Definition and (a b : Bool) : Bool
|
||||
:= ¬ (a ⇒ ¬ b).
|
||||
definition and (a b : Bool) : Bool
|
||||
:= ¬ (a ⇒ ¬ b)
|
||||
|
||||
Infixr 35 && : and.
|
||||
Infixr 35 /\ : and.
|
||||
Infixr 35 ∧ : and.
|
||||
infixr 35 && : and
|
||||
infixr 35 /\ : and
|
||||
infixr 35 ∧ : and
|
||||
|
||||
-- Forall is a macro for the identifier forall, we use that
|
||||
-- because the Lean parser has the builtin syntax sugar:
|
||||
-- forall x : T, P x
|
||||
-- for
|
||||
-- (forall T (fun x : T, P x))
|
||||
Definition Forall (A : TypeU) (P : A → Bool) : Bool
|
||||
:= P == (λ x : A, true).
|
||||
definition Forall (A : TypeU) (P : A → Bool) : Bool
|
||||
:= P == (λ x : A, true)
|
||||
|
||||
Definition Exists (A : TypeU) (P : A → Bool) : Bool
|
||||
:= ¬ (Forall A (λ x : A, ¬ (P x))).
|
||||
definition Exists (A : TypeU) (P : A → Bool) : Bool
|
||||
:= ¬ (Forall A (λ x : A, ¬ (P x)))
|
||||
|
||||
Definition eq {A : TypeU} (a b : A) : Bool
|
||||
:= a == b.
|
||||
definition eq {A : TypeU} (a b : A) : Bool
|
||||
:= a == b
|
||||
|
||||
Infix 50 = : eq.
|
||||
infix 50 = : eq
|
||||
|
||||
Definition neq {A : TypeU} (a b : A) : Bool
|
||||
:= ¬ (a == b).
|
||||
definition neq {A : TypeU} (a b : A) : Bool
|
||||
:= ¬ (a == b)
|
||||
|
||||
Infix 50 ≠ : neq.
|
||||
infix 50 ≠ : neq
|
||||
|
||||
Axiom MP {a b : Bool} (H1 : a ⇒ b) (H2 : a) : b.
|
||||
axiom MP {a b : Bool} (H1 : a ⇒ b) (H2 : a) : b
|
||||
|
||||
Axiom Discharge {a b : Bool} (H : a → b) : a ⇒ b.
|
||||
axiom Discharge {a b : Bool} (H : a → b) : a ⇒ b
|
||||
|
||||
Axiom Case (P : Bool → Bool) (H1 : P true) (H2 : P false) (a : Bool) : P a.
|
||||
axiom Case (P : Bool → Bool) (H1 : P true) (H2 : P false) (a : Bool) : P a
|
||||
|
||||
Axiom Refl {A : TypeU} (a : A) : a == a.
|
||||
axiom Refl {A : TypeU} (a : A) : a == a
|
||||
|
||||
Axiom Subst {A : TypeU} {a b : A} {P : A → Bool} (H1 : P a) (H2 : a == b) : P b.
|
||||
axiom Subst {A : TypeU} {a b : A} {P : A → Bool} (H1 : P a) (H2 : a == b) : P b
|
||||
|
||||
Definition SubstP {A : TypeU} {a b : A} (P : A → Bool) (H1 : P a) (H2 : a == b) : P b
|
||||
:= Subst H1 H2.
|
||||
definition SubstP {A : TypeU} {a b : A} (P : A → Bool) (H1 : P a) (H2 : a == b) : P b
|
||||
:= Subst H1 H2
|
||||
|
||||
Axiom Eta {A : TypeU} {B : A → TypeU} (f : Π x : A, B x) : (λ x : A, f x) == f.
|
||||
axiom Eta {A : TypeU} {B : A → TypeU} (f : Π x : A, B x) : (λ x : A, f x) == f
|
||||
|
||||
Axiom ImpAntisym {a b : Bool} (H1 : a ⇒ b) (H2 : b ⇒ a) : a == b.
|
||||
axiom ImpAntisym {a b : Bool} (H1 : a ⇒ b) (H2 : b ⇒ a) : a == b
|
||||
|
||||
Axiom Abst {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} (H : Π x : A, f x == g x) : f == g.
|
||||
axiom Abst {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} (H : Π x : A, f x == g x) : f == g
|
||||
|
||||
Axiom HSymm {A B : TypeU} {a : A} {b : B} (H : a == b) : b == a.
|
||||
axiom HSymm {A B : TypeU} {a : A} {b : B} (H : a == b) : b == a
|
||||
|
||||
Axiom HTrans {A B C : TypeU} {a : A} {b : B} {c : C} (H1 : a == b) (H2 : b == c) : a == c.
|
||||
axiom HTrans {A B C : TypeU} {a : A} {b : B} {c : C} (H1 : a == b) (H2 : b == c) : a == c
|
||||
|
||||
Theorem Trivial : true
|
||||
:= Refl true.
|
||||
theorem Trivial : true
|
||||
:= Refl true
|
||||
|
||||
Theorem TrueNeFalse : not (true == false)
|
||||
:= Trivial.
|
||||
theorem TrueNeFalse : not (true == false)
|
||||
:= Trivial
|
||||
|
||||
Theorem EM (a : Bool) : a ∨ ¬ a
|
||||
:= Case (λ x, x ∨ ¬ x) Trivial Trivial a.
|
||||
theorem EM (a : Bool) : a ∨ ¬ a
|
||||
:= Case (λ x, x ∨ ¬ x) Trivial Trivial a
|
||||
|
||||
Theorem FalseElim (a : Bool) (H : false) : a
|
||||
:= Case (λ x, x) Trivial H a.
|
||||
theorem FalseElim (a : Bool) (H : false) : a
|
||||
:= Case (λ x, x) Trivial H a
|
||||
|
||||
Theorem Absurd {a : Bool} (H1 : a) (H2 : ¬ a) : false
|
||||
:= MP H2 H1.
|
||||
theorem Absurd {a : Bool} (H1 : a) (H2 : ¬ a) : false
|
||||
:= MP H2 H1
|
||||
|
||||
Theorem EqMP {a b : Bool} (H1 : a == b) (H2 : a) : b
|
||||
:= Subst H2 H1.
|
||||
theorem EqMP {a b : Bool} (H1 : a == b) (H2 : a) : b
|
||||
:= Subst H2 H1
|
||||
|
||||
-- Assume is a 'macro' that expands into a Discharge
|
||||
|
||||
Theorem ImpTrans {a b c : Bool} (H1 : a ⇒ b) (H2 : b ⇒ c) : a ⇒ c
|
||||
:= Assume Ha, MP H2 (MP H1 Ha).
|
||||
theorem ImpTrans {a b c : Bool} (H1 : a ⇒ b) (H2 : b ⇒ c) : a ⇒ c
|
||||
:= Assume Ha, MP H2 (MP H1 Ha)
|
||||
|
||||
Theorem ImpEqTrans {a b c : Bool} (H1 : a ⇒ b) (H2 : b == c) : a ⇒ c
|
||||
:= Assume Ha, EqMP H2 (MP H1 Ha).
|
||||
theorem ImpEqTrans {a b c : Bool} (H1 : a ⇒ b) (H2 : b == c) : a ⇒ c
|
||||
:= Assume Ha, EqMP H2 (MP H1 Ha)
|
||||
|
||||
Theorem EqImpTrans {a b c : Bool} (H1 : a == b) (H2 : b ⇒ c) : a ⇒ c
|
||||
:= Assume Ha, MP H2 (EqMP H1 Ha).
|
||||
theorem EqImpTrans {a b c : Bool} (H1 : a == b) (H2 : b ⇒ c) : a ⇒ c
|
||||
:= Assume Ha, MP H2 (EqMP H1 Ha)
|
||||
|
||||
Theorem DoubleNeg (a : Bool) : (¬ ¬ a) == a
|
||||
:= Case (λ x, (¬ ¬ x) == x) Trivial Trivial a.
|
||||
theorem DoubleNeg (a : Bool) : (¬ ¬ a) == a
|
||||
:= Case (λ x, (¬ ¬ x) == x) Trivial Trivial a
|
||||
|
||||
Theorem DoubleNegElim {a : Bool} (H : ¬ ¬ a) : a
|
||||
:= EqMP (DoubleNeg a) H.
|
||||
theorem DoubleNegElim {a : Bool} (H : ¬ ¬ a) : a
|
||||
:= EqMP (DoubleNeg a) H
|
||||
|
||||
Theorem MT {a b : Bool} (H1 : a ⇒ b) (H2 : ¬ b) : ¬ a
|
||||
:= Assume H : a, Absurd (MP H1 H) H2.
|
||||
theorem MT {a b : Bool} (H1 : a ⇒ b) (H2 : ¬ b) : ¬ a
|
||||
:= Assume H : a, Absurd (MP H1 H) H2
|
||||
|
||||
Theorem Contrapos {a b : Bool} (H : a ⇒ b) : ¬ b ⇒ ¬ a
|
||||
:= Assume H1 : ¬ b, MT H H1.
|
||||
theorem Contrapos {a b : Bool} (H : a ⇒ b) : ¬ b ⇒ ¬ a
|
||||
:= Assume H1 : ¬ b, MT H H1
|
||||
|
||||
Theorem AbsurdElim {a : Bool} (b : Bool) (H1 : a) (H2 : ¬ a) : b
|
||||
:= FalseElim b (Absurd H1 H2).
|
||||
theorem AbsurdElim {a : Bool} (b : Bool) (H1 : a) (H2 : ¬ a) : b
|
||||
:= FalseElim b (Absurd H1 H2)
|
||||
|
||||
Theorem NotImp1 {a b : Bool} (H : ¬ (a ⇒ b)) : a
|
||||
theorem NotImp1 {a b : Bool} (H : ¬ (a ⇒ b)) : a
|
||||
:= DoubleNegElim
|
||||
(have ¬ ¬ a :
|
||||
Assume H1 : ¬ a, Absurd (have a ⇒ b : Assume H2 : a, AbsurdElim b H2 H1)
|
||||
(have ¬ (a ⇒ b) : H)).
|
||||
(have ¬ (a ⇒ b) : H))
|
||||
|
||||
Theorem NotImp2 {a b : Bool} (H : ¬ (a ⇒ b)) : ¬ b
|
||||
theorem NotImp2 {a b : Bool} (H : ¬ (a ⇒ b)) : ¬ b
|
||||
:= Assume H1 : b, Absurd (have a ⇒ b : Assume H2 : a, H1)
|
||||
(have ¬ (a ⇒ b) : H).
|
||||
(have ¬ (a ⇒ b) : H)
|
||||
|
||||
-- Remark: conjunction is defined as ¬ (a ⇒ ¬ b) in Lean
|
||||
|
||||
Theorem Conj {a b : Bool} (H1 : a) (H2 : b) : a ∧ b
|
||||
:= Assume H : a ⇒ ¬ b, Absurd H2 (MP H H1).
|
||||
theorem Conj {a b : Bool} (H1 : a) (H2 : b) : a ∧ b
|
||||
:= Assume H : a ⇒ ¬ b, Absurd H2 (MP H H1)
|
||||
|
||||
Theorem Conjunct1 {a b : Bool} (H : a ∧ b) : a
|
||||
:= NotImp1 H.
|
||||
theorem Conjunct1 {a b : Bool} (H : a ∧ b) : a
|
||||
:= NotImp1 H
|
||||
|
||||
Theorem Conjunct2 {a b : Bool} (H : a ∧ b) : b
|
||||
:= DoubleNegElim (NotImp2 H).
|
||||
theorem Conjunct2 {a b : Bool} (H : a ∧ b) : b
|
||||
:= DoubleNegElim (NotImp2 H)
|
||||
|
||||
-- Remark: disjunction is defined as ¬ a ⇒ b in Lean
|
||||
|
||||
Theorem Disj1 {a : Bool} (H : a) (b : Bool) : a ∨ b
|
||||
:= Assume H1 : ¬ a, AbsurdElim b H H1.
|
||||
theorem Disj1 {a : Bool} (H : a) (b : Bool) : a ∨ b
|
||||
:= Assume H1 : ¬ a, AbsurdElim b H H1
|
||||
|
||||
Theorem Disj2 {b : Bool} (a : Bool) (H : b) : a ∨ b
|
||||
:= Assume H1 : ¬ a, H.
|
||||
theorem Disj2 {b : Bool} (a : Bool) (H : b) : a ∨ b
|
||||
:= Assume H1 : ¬ a, H
|
||||
|
||||
Theorem ArrowToImplies {a b : Bool} (H : a → b) : a ⇒ b
|
||||
:= Assume H1 : a, H H1.
|
||||
theorem ArrowToImplies {a b : Bool} (H : a → b) : a ⇒ b
|
||||
:= Assume H1 : a, H H1
|
||||
|
||||
Theorem Resolve1 {a b : Bool} (H1 : a ∨ b) (H2 : ¬ a) : b
|
||||
:= MP H1 H2.
|
||||
theorem Resolve1 {a b : Bool} (H1 : a ∨ b) (H2 : ¬ a) : b
|
||||
:= MP H1 H2
|
||||
|
||||
Theorem DisjCases {a b c : Bool} (H1 : a ∨ b) (H2 : a → c) (H3 : b → c) : c
|
||||
theorem DisjCases {a b c : Bool} (H1 : a ∨ b) (H2 : a → c) (H3 : b → c) : c
|
||||
:= DoubleNegElim
|
||||
(Assume H : ¬ c,
|
||||
Absurd (have c : H3 (have b : Resolve1 H1 (have ¬ a : (MT (ArrowToImplies H2) H))))
|
||||
H)
|
||||
|
||||
Theorem Refute {a : Bool} (H : ¬ a → false) : a
|
||||
:= DisjCases (EM a) (λ H1 : a, H1) (λ H1 : ¬ a, FalseElim a (H H1)).
|
||||
theorem Refute {a : Bool} (H : ¬ a → false) : a
|
||||
:= DisjCases (EM a) (λ H1 : a, H1) (λ H1 : ¬ a, FalseElim a (H H1))
|
||||
|
||||
Theorem Symm {A : TypeU} {a b : A} (H : a == b) : b == a
|
||||
:= Subst (Refl a) H.
|
||||
theorem Symm {A : TypeU} {a b : A} (H : a == b) : b == a
|
||||
:= Subst (Refl a) H
|
||||
|
||||
Theorem NeSymm {A : TypeU} {a b : A} (H : a ≠ b) : b ≠ a
|
||||
:= Assume H1 : b = a, MP H (Symm H1).
|
||||
theorem NeSymm {A : TypeU} {a b : A} (H : a ≠ b) : b ≠ a
|
||||
:= Assume H1 : b = a, MP H (Symm H1)
|
||||
|
||||
Theorem EqNeTrans {A : TypeU} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c
|
||||
:= Subst H2 (Symm H1).
|
||||
theorem EqNeTrans {A : TypeU} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c
|
||||
:= Subst H2 (Symm H1)
|
||||
|
||||
Theorem NeEqTrans {A : TypeU} {a b c : A} (H1 : a ≠ b) (H2 : b = c) : a ≠ c
|
||||
:= Subst H1 H2.
|
||||
theorem NeEqTrans {A : TypeU} {a b c : A} (H1 : a ≠ b) (H2 : b = c) : a ≠ c
|
||||
:= Subst H1 H2
|
||||
|
||||
Theorem Trans {A : TypeU} {a b c : A} (H1 : a == b) (H2 : b == c) : a == c
|
||||
:= Subst H1 H2.
|
||||
theorem Trans {A : TypeU} {a b c : A} (H1 : a == b) (H2 : b == c) : a == c
|
||||
:= Subst H1 H2
|
||||
|
||||
Theorem EqTElim {a : Bool} (H : a == true) : a
|
||||
:= EqMP (Symm H) Trivial.
|
||||
theorem EqTElim {a : Bool} (H : a == true) : a
|
||||
:= EqMP (Symm H) Trivial
|
||||
|
||||
Theorem EqTIntro {a : Bool} (H : a) : a == true
|
||||
theorem EqTIntro {a : Bool} (H : a) : a == true
|
||||
:= ImpAntisym (Assume H1 : a, Trivial)
|
||||
(Assume H2 : true, H).
|
||||
(Assume H2 : true, H)
|
||||
|
||||
Theorem Congr1 {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} (a : A) (H : f == g) : f a == g a
|
||||
:= SubstP (fun h : (Π x : A, B x), f a == h a) (Refl (f a)) H.
|
||||
theorem Congr1 {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} (a : A) (H : f == g) : f a == g a
|
||||
:= SubstP (fun h : (Π x : A, B x), f a == h a) (Refl (f a)) H
|
||||
|
||||
-- Remark: we must use heterogeneous equality in the following theorem because the types of (f a) and (f b)
|
||||
-- are not "definitionally equal". They are (B a) and (B b).
|
||||
-- They are provably equal, we just have to apply Congr1.
|
||||
-- are not "definitionally equal" They are (B a) and (B b)
|
||||
-- They are provably equal, we just have to apply Congr1
|
||||
|
||||
Theorem Congr2 {A : TypeU} {B : A → TypeU} {a b : A} (f : Π x : A, B x) (H : a == b) : f a == f b
|
||||
:= SubstP (fun x : A, f a == f x) (Refl (f a)) H.
|
||||
theorem Congr2 {A : TypeU} {B : A → TypeU} {a b : A} (f : Π x : A, B x) (H : a == b) : f a == f b
|
||||
:= SubstP (fun x : A, f a == f x) (Refl (f a)) H
|
||||
|
||||
-- Remark: like the previous theorem we use heterogeneous equality. We cannot use Trans theorem
|
||||
-- because the types are not definitionally equal.
|
||||
-- Remark: like the previous theorem we use heterogeneous equality We cannot use Trans theorem
|
||||
-- because the types are not definitionally equal
|
||||
|
||||
Theorem Congr {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} {a b : A} (H1 : f == g) (H2 : a == b) : f a == g b
|
||||
:= HTrans (Congr2 f H2) (Congr1 b H1).
|
||||
theorem Congr {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} {a b : A} (H1 : f == g) (H2 : a == b) : f a == g b
|
||||
:= HTrans (Congr2 f H2) (Congr1 b H1)
|
||||
|
||||
Theorem ForallElim {A : TypeU} {P : A → Bool} (H : Forall A P) (a : A) : P a
|
||||
:= EqTElim (Congr1 a H).
|
||||
theorem ForallElim {A : TypeU} {P : A → Bool} (H : Forall A P) (a : A) : P a
|
||||
:= EqTElim (Congr1 a H)
|
||||
|
||||
Theorem ForallIntro {A : TypeU} {P : A → Bool} (H : Π x : A, P x) : Forall A P
|
||||
theorem ForallIntro {A : TypeU} {P : A → Bool} (H : Π x : A, P x) : Forall A P
|
||||
:= Trans (Symm (Eta P))
|
||||
(Abst (λ x, EqTIntro (H x))).
|
||||
(Abst (λ x, EqTIntro (H x)))
|
||||
|
||||
-- Remark: the existential is defined as (¬ (forall x : A, ¬ P x))
|
||||
|
||||
Theorem ExistsElim {A : TypeU} {P : A → Bool} {B : Bool} (H1 : Exists A P) (H2 : Π (a : A) (H : P a), B) : B
|
||||
theorem ExistsElim {A : TypeU} {P : A → Bool} {B : Bool} (H1 : Exists A P) (H2 : Π (a : A) (H : P a), B) : B
|
||||
:= Refute (λ R : ¬ B,
|
||||
Absurd (ForallIntro (λ a : A, MT (Assume H : P a, H2 a H) R))
|
||||
H1).
|
||||
H1)
|
||||
|
||||
Theorem ExistsIntro {A : TypeU} {P : A → Bool} (a : A) (H : P a) : Exists A P
|
||||
theorem ExistsIntro {A : TypeU} {P : A → Bool} (a : A) (H : P a) : Exists A P
|
||||
:= Assume H1 : (∀ x : A, ¬ P x),
|
||||
Absurd H (ForallElim H1 a).
|
||||
Absurd H (ForallElim H1 a)
|
||||
|
||||
-- At this point, we have proved the theorems we need using the
|
||||
-- definitions of forall, exists, and, or, =>, not. We mark (some of)
|
||||
-- them as opaque. Opaque definitions improve performance, and
|
||||
-- effectiveness of Lean's elaborator.
|
||||
-- definitions of forall, exists, and, or, =>, not We mark (some of)
|
||||
-- them as opaque Opaque definitions improve performance, and
|
||||
-- effectiveness of Lean's elaborator
|
||||
|
||||
SetOpaque implies true.
|
||||
SetOpaque not true.
|
||||
SetOpaque or true.
|
||||
SetOpaque and true.
|
||||
SetOpaque forall true.
|
||||
setopaque implies true
|
||||
setopaque not true
|
||||
setopaque or true
|
||||
setopaque and true
|
||||
setopaque forall true
|
||||
|
||||
Theorem OrComm (a b : Bool) : (a ∨ b) == (b ∨ a)
|
||||
theorem OrComm (a b : Bool) : (a ∨ b) == (b ∨ a)
|
||||
:= ImpAntisym (Assume H, DisjCases H (λ H1, Disj2 b H1) (λ H2, Disj1 H2 a))
|
||||
(Assume H, DisjCases H (λ H1, Disj2 a H1) (λ H2, Disj1 H2 b)).
|
||||
(Assume H, DisjCases H (λ H1, Disj2 a H1) (λ H2, Disj1 H2 b))
|
||||
|
||||
|
||||
Theorem OrAssoc (a b c : Bool) : ((a ∨ b) ∨ c) == (a ∨ (b ∨ c))
|
||||
theorem OrAssoc (a b c : Bool) : ((a ∨ b) ∨ c) == (a ∨ (b ∨ c))
|
||||
:= ImpAntisym (Assume H : (a ∨ b) ∨ c,
|
||||
DisjCases H (λ H1 : a ∨ b, DisjCases H1 (λ Ha : a, Disj1 Ha (b ∨ c)) (λ Hb : b, Disj2 a (Disj1 Hb c)))
|
||||
(λ Hc : c, Disj2 a (Disj2 b Hc)))
|
||||
(Assume H : a ∨ (b ∨ c),
|
||||
DisjCases H (λ Ha : a, (Disj1 (Disj1 Ha b) c))
|
||||
(λ H1 : b ∨ c, DisjCases H1 (λ Hb : b, Disj1 (Disj2 a Hb) c)
|
||||
(λ Hc : c, Disj2 (a ∨ b) Hc))).
|
||||
(λ Hc : c, Disj2 (a ∨ b) Hc)))
|
||||
|
||||
Theorem OrId (a : Bool) : (a ∨ a) == a
|
||||
theorem OrId (a : Bool) : (a ∨ a) == a
|
||||
:= ImpAntisym (Assume H, DisjCases H (λ H1, H1) (λ H2, H2))
|
||||
(Assume H, Disj1 H a).
|
||||
(Assume H, Disj1 H a)
|
||||
|
||||
Theorem OrRhsFalse (a : Bool) : (a ∨ false) == a
|
||||
theorem OrRhsFalse (a : Bool) : (a ∨ false) == a
|
||||
:= ImpAntisym (Assume H, DisjCases H (λ H1, H1) (λ H2, FalseElim a H2))
|
||||
(Assume H, Disj1 H false).
|
||||
(Assume H, Disj1 H false)
|
||||
|
||||
Theorem OrLhsFalse (a : Bool) : (false ∨ a) == a
|
||||
:= Trans (OrComm false a) (OrRhsFalse a).
|
||||
theorem OrLhsFalse (a : Bool) : (false ∨ a) == a
|
||||
:= Trans (OrComm false a) (OrRhsFalse a)
|
||||
|
||||
Theorem OrLhsTrue (a : Bool) : (true ∨ a) == true
|
||||
:= EqTIntro (Case (λ x : Bool, true ∨ x) Trivial Trivial a).
|
||||
theorem OrLhsTrue (a : Bool) : (true ∨ a) == true
|
||||
:= EqTIntro (Case (λ x : Bool, true ∨ x) Trivial Trivial a)
|
||||
|
||||
Theorem OrRhsTrue (a : Bool) : (a ∨ true) == true
|
||||
:= Trans (OrComm a true) (OrLhsTrue a).
|
||||
theorem OrRhsTrue (a : Bool) : (a ∨ true) == true
|
||||
:= Trans (OrComm a true) (OrLhsTrue a)
|
||||
|
||||
Theorem OrAnotA (a : Bool) : (a ∨ ¬ a) == true
|
||||
:= EqTIntro (EM a).
|
||||
theorem OrAnotA (a : Bool) : (a ∨ ¬ a) == true
|
||||
:= EqTIntro (EM a)
|
||||
|
||||
Theorem AndComm (a b : Bool) : (a ∧ b) == (b ∧ a)
|
||||
theorem AndComm (a b : Bool) : (a ∧ b) == (b ∧ a)
|
||||
:= ImpAntisym (Assume H, Conj (Conjunct2 H) (Conjunct1 H))
|
||||
(Assume H, Conj (Conjunct2 H) (Conjunct1 H)).
|
||||
(Assume H, Conj (Conjunct2 H) (Conjunct1 H))
|
||||
|
||||
Theorem AndId (a : Bool) : (a ∧ a) == a
|
||||
theorem AndId (a : Bool) : (a ∧ a) == a
|
||||
:= ImpAntisym (Assume H, Conjunct1 H)
|
||||
(Assume H, Conj H H).
|
||||
(Assume H, Conj H H)
|
||||
|
||||
Theorem AndAssoc (a b c : Bool) : ((a ∧ b) ∧ c) == (a ∧ (b ∧ c))
|
||||
theorem AndAssoc (a b c : Bool) : ((a ∧ b) ∧ c) == (a ∧ (b ∧ c))
|
||||
:= ImpAntisym (Assume H, Conj (Conjunct1 (Conjunct1 H)) (Conj (Conjunct2 (Conjunct1 H)) (Conjunct2 H)))
|
||||
(Assume H, Conj (Conj (Conjunct1 H) (Conjunct1 (Conjunct2 H))) (Conjunct2 (Conjunct2 H))).
|
||||
(Assume H, Conj (Conj (Conjunct1 H) (Conjunct1 (Conjunct2 H))) (Conjunct2 (Conjunct2 H)))
|
||||
|
||||
Theorem AndRhsTrue (a : Bool) : (a ∧ true) == a
|
||||
theorem AndRhsTrue (a : Bool) : (a ∧ true) == a
|
||||
:= ImpAntisym (Assume H : a ∧ true, Conjunct1 H)
|
||||
(Assume H : a, Conj H Trivial).
|
||||
(Assume H : a, Conj H Trivial)
|
||||
|
||||
Theorem AndLhsTrue (a : Bool) : (true ∧ a) == a
|
||||
:= Trans (AndComm true a) (AndRhsTrue a).
|
||||
theorem AndLhsTrue (a : Bool) : (true ∧ a) == a
|
||||
:= Trans (AndComm true a) (AndRhsTrue a)
|
||||
|
||||
Theorem AndRhsFalse (a : Bool) : (a ∧ false) == false
|
||||
theorem AndRhsFalse (a : Bool) : (a ∧ false) == false
|
||||
:= ImpAntisym (Assume H, Conjunct2 H)
|
||||
(Assume H, FalseElim (a ∧ false) H).
|
||||
(Assume H, FalseElim (a ∧ false) H)
|
||||
|
||||
Theorem AndLhsFalse (a : Bool) : (false ∧ a) == false
|
||||
:= Trans (AndComm false a) (AndRhsFalse a).
|
||||
theorem AndLhsFalse (a : Bool) : (false ∧ a) == false
|
||||
:= Trans (AndComm false a) (AndRhsFalse a)
|
||||
|
||||
Theorem AndAnotA (a : Bool) : (a ∧ ¬ a) == false
|
||||
theorem AndAnotA (a : Bool) : (a ∧ ¬ a) == false
|
||||
:= ImpAntisym (Assume H, Absurd (Conjunct1 H) (Conjunct2 H))
|
||||
(Assume H, FalseElim (a ∧ ¬ a) H).
|
||||
(Assume H, FalseElim (a ∧ ¬ a) H)
|
||||
|
||||
Theorem NotTrue : (¬ true) == false
|
||||
theorem NotTrue : (¬ true) == false
|
||||
:= Trivial
|
||||
|
||||
Theorem NotFalse : (¬ false) == true
|
||||
theorem NotFalse : (¬ false) == true
|
||||
:= Trivial
|
||||
|
||||
Theorem NotAnd (a b : Bool) : (¬ (a ∧ b)) == (¬ a ∨ ¬ b)
|
||||
theorem NotAnd (a b : Bool) : (¬ (a ∧ b)) == (¬ a ∨ ¬ b)
|
||||
:= Case (λ x, (¬ (x ∧ b)) == (¬ x ∨ ¬ b))
|
||||
(Case (λ y, (¬ (true ∧ y)) == (¬ true ∨ ¬ y)) Trivial Trivial b)
|
||||
(Case (λ y, (¬ (false ∧ y)) == (¬ false ∨ ¬ y)) Trivial Trivial b)
|
||||
a
|
||||
|
||||
Theorem NotAndElim {a b : Bool} (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b
|
||||
:= EqMP (NotAnd a b) H.
|
||||
theorem NotAndElim {a b : Bool} (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b
|
||||
:= EqMP (NotAnd a b) H
|
||||
|
||||
Theorem NotOr (a b : Bool) : (¬ (a ∨ b)) == (¬ a ∧ ¬ b)
|
||||
theorem NotOr (a b : Bool) : (¬ (a ∨ b)) == (¬ a ∧ ¬ b)
|
||||
:= Case (λ x, (¬ (x ∨ b)) == (¬ x ∧ ¬ b))
|
||||
(Case (λ y, (¬ (true ∨ y)) == (¬ true ∧ ¬ y)) Trivial Trivial b)
|
||||
(Case (λ y, (¬ (false ∨ y)) == (¬ false ∧ ¬ y)) Trivial Trivial b)
|
||||
a
|
||||
|
||||
Theorem NotOrElim {a b : Bool} (H : ¬ (a ∨ b)) : ¬ a ∧ ¬ b
|
||||
:= EqMP (NotOr a b) H.
|
||||
theorem NotOrElim {a b : Bool} (H : ¬ (a ∨ b)) : ¬ a ∧ ¬ b
|
||||
:= EqMP (NotOr a b) H
|
||||
|
||||
Theorem NotEq (a b : Bool) : (¬ (a == b)) == ((¬ a) == b)
|
||||
theorem NotEq (a b : Bool) : (¬ (a == b)) == ((¬ a) == b)
|
||||
:= Case (λ x, (¬ (x == b)) == ((¬ x) == b))
|
||||
(Case (λ y, (¬ (true == y)) == ((¬ true) == y)) Trivial Trivial b)
|
||||
(Case (λ y, (¬ (false == y)) == ((¬ false) == y)) Trivial Trivial b)
|
||||
a
|
||||
|
||||
Theorem NotEqElim {a b : Bool} (H : ¬ (a == b)) : (¬ a) == b
|
||||
:= EqMP (NotEq a b) H.
|
||||
theorem NotEqElim {a b : Bool} (H : ¬ (a == b)) : (¬ a) == b
|
||||
:= EqMP (NotEq a b) H
|
||||
|
||||
Theorem NotImp (a b : Bool) : (¬ (a ⇒ b)) == (a ∧ ¬ b)
|
||||
theorem NotImp (a b : Bool) : (¬ (a ⇒ b)) == (a ∧ ¬ b)
|
||||
:= Case (λ x, (¬ (x ⇒ b)) == (x ∧ ¬ b))
|
||||
(Case (λ y, (¬ (true ⇒ y)) == (true ∧ ¬ y)) Trivial Trivial b)
|
||||
(Case (λ y, (¬ (false ⇒ y)) == (false ∧ ¬ y)) Trivial Trivial b)
|
||||
a
|
||||
|
||||
Theorem NotImpElim {a b : Bool} (H : ¬ (a ⇒ b)) : a ∧ ¬ b
|
||||
:= EqMP (NotImp a b) H.
|
||||
theorem NotImpElim {a b : Bool} (H : ¬ (a ⇒ b)) : a ∧ ¬ b
|
||||
:= EqMP (NotImp a b) H
|
||||
|
||||
Theorem NotCongr {a b : Bool} (H : a == b) : (¬ a) == (¬ b)
|
||||
:= Congr2 not H.
|
||||
theorem NotCongr {a b : Bool} (H : a == b) : (¬ a) == (¬ b)
|
||||
:= Congr2 not H
|
||||
|
||||
Theorem ForallEqIntro {A : (Type U)} {P Q : A → Bool} (H : Pi x : A, P x == Q x) : (∀ x : A, P x) == (∀ x : A, Q x)
|
||||
:= Congr2 (Forall A) (Abst H).
|
||||
theorem ForallEqIntro {A : (Type U)} {P Q : A → Bool} (H : Pi x : A, P x == Q x) : (∀ x : A, P x) == (∀ x : A, Q x)
|
||||
:= Congr2 (Forall A) (Abst H)
|
||||
|
||||
Theorem ExistsEqIntro {A : (Type U)} {P Q : A → Bool} (H : Pi x : A, P x == Q x) : (∃ x : A, P x) == (∃ x : A, Q x)
|
||||
:= Congr2 (Exists A) (Abst H).
|
||||
theorem ExistsEqIntro {A : (Type U)} {P Q : A → Bool} (H : Pi x : A, P x == Q x) : (∃ x : A, P x) == (∃ x : A, Q x)
|
||||
:= Congr2 (Exists A) (Abst H)
|
||||
|
||||
Theorem NotForall (A : (Type U)) (P : A → Bool) : (¬ (∀ x : A, P x)) == (∃ x : A, ¬ P x)
|
||||
theorem NotForall (A : (Type U)) (P : A → Bool) : (¬ (∀ x : A, P x)) == (∃ x : A, ¬ P x)
|
||||
:= let L1 : (¬ ∀ x : A, ¬ ¬ P x) == (∃ x : A, ¬ P x) := Refl (∃ x : A, ¬ P x),
|
||||
L2 : (¬ ∀ x : A, P x) == (¬ ∀ x : A, ¬ ¬ P x) :=
|
||||
NotCongr (ForallEqIntro (λ x : A, (Symm (DoubleNeg (P x)))))
|
||||
in Trans L2 L1.
|
||||
in Trans L2 L1
|
||||
|
||||
Theorem NotForallElim {A : (Type U)} {P : A → Bool} (H : ¬ (∀ x : A, P x)) : ∃ x : A, ¬ P x
|
||||
:= EqMP (NotForall A P) H.
|
||||
theorem NotForallElim {A : (Type U)} {P : A → Bool} (H : ¬ (∀ x : A, P x)) : ∃ x : A, ¬ P x
|
||||
:= EqMP (NotForall A P) H
|
||||
|
||||
Theorem NotExists (A : (Type U)) (P : A → Bool) : (¬ ∃ x : A, P x) == (∀ x : A, ¬ P x)
|
||||
theorem NotExists (A : (Type U)) (P : A → Bool) : (¬ ∃ x : A, P x) == (∀ x : A, ¬ P x)
|
||||
:= let L1 : (¬ ∃ x : A, P x) == (¬ ¬ ∀ x : A, ¬ P x) := Refl (¬ ∃ x : A, P x),
|
||||
L2 : (¬ ¬ ∀ x : A, ¬ P x) == (∀ x : A, ¬ P x) := DoubleNeg (∀ x : A, ¬ P x)
|
||||
in Trans L1 L2.
|
||||
in Trans L1 L2
|
||||
|
||||
Theorem NotExistsElim {A : (Type U)} {P : A → Bool} (H : ¬ ∃ x : A, P x) : ∀ x : A, ¬ P x
|
||||
:= EqMP (NotExists A P) H.
|
||||
theorem NotExistsElim {A : (Type U)} {P : A → Bool} (H : ¬ ∃ x : A, P x) : ∀ x : A, ¬ P x
|
||||
:= EqMP (NotExists A P) H
|
||||
|
||||
Theorem UnfoldExists1 {A : TypeU} {P : A → Bool} (a : A) (H : ∃ x : A, P x) : P a ∨ (∃ x : A, x ≠ a ∧ P x)
|
||||
theorem UnfoldExists1 {A : TypeU} {P : A → Bool} (a : A) (H : ∃ x : A, P x) : P a ∨ (∃ x : A, x ≠ a ∧ P x)
|
||||
:= ExistsElim H
|
||||
(λ (w : A) (H1 : P w),
|
||||
DisjCases (EM (w = a))
|
||||
(λ Heq : w = a, Disj1 (Subst H1 Heq) (∃ x : A, x ≠ a ∧ P x))
|
||||
(λ Hne : w ≠ a, Disj2 (P a) (ExistsIntro w (Conj Hne H1)))).
|
||||
(λ Hne : w ≠ a, Disj2 (P a) (ExistsIntro w (Conj Hne H1))))
|
||||
|
||||
Theorem UnfoldExists2 {A : TypeU} {P : A → Bool} (a : A) (H : P a ∨ (∃ x : A, x ≠ a ∧ P x)) : ∃ x : A, P x
|
||||
theorem UnfoldExists2 {A : TypeU} {P : A → Bool} (a : A) (H : P a ∨ (∃ x : A, x ≠ a ∧ P x)) : ∃ x : A, P x
|
||||
:= DisjCases H
|
||||
(λ H1 : P a, ExistsIntro a H1)
|
||||
(λ H2 : (∃ x : A, x ≠ a ∧ P x),
|
||||
ExistsElim H2
|
||||
(λ (w : A) (Hw : w ≠ a ∧ P w),
|
||||
ExistsIntro w (Conjunct2 Hw))).
|
||||
ExistsIntro w (Conjunct2 Hw)))
|
||||
|
||||
Theorem UnfoldExists {A : TypeU} (P : A → Bool) (a : A) : (∃ x : A, P x) = (P a ∨ (∃ x : A, x ≠ a ∧ P x))
|
||||
theorem UnfoldExists {A : TypeU} (P : A → Bool) (a : A) : (∃ x : A, P x) = (P a ∨ (∃ x : A, x ≠ a ∧ P x))
|
||||
:= ImpAntisym (Assume H : (∃ x : A, P x), UnfoldExists1 a H)
|
||||
(Assume H : (P a ∨ (∃ x : A, x ≠ a ∧ P x)), UnfoldExists2 a H).
|
||||
(Assume H : (P a ∨ (∃ x : A, x ≠ a ∧ P x)), UnfoldExists2 a H)
|
||||
|
||||
SetOpaque exists true.
|
||||
setopaque exists true
|
||||
|
|
Binary file not shown.
Binary file not shown.
Binary file not shown.
Binary file not shown.
Binary file not shown.
Binary file not shown.
|
@ -1,19 +1,19 @@
|
|||
Import Real.
|
||||
import Real.
|
||||
|
||||
Variable exp : Real → Real.
|
||||
Variable log : Real → Real.
|
||||
Variable pi : Real.
|
||||
Alias π : pi.
|
||||
variable exp : Real → Real.
|
||||
variable log : Real → Real.
|
||||
variable pi : Real.
|
||||
alias π : pi.
|
||||
|
||||
Variable sin : Real → Real.
|
||||
Definition cos x := sin (x - π / 2).
|
||||
Definition tan x := sin x / cos x.
|
||||
Definition cot x := cos x / sin x.
|
||||
Definition sec x := 1 / cos x.
|
||||
Definition csc x := 1 / sin x.
|
||||
Definition sinh x := (1 - exp (-2 * x)) / (2 * exp (- x)).
|
||||
Definition cosh x := (1 + exp (-2 * x)) / (2 * exp (- x)).
|
||||
Definition tanh x := sinh x / cosh x.
|
||||
Definition coth x := cosh x / sinh x.
|
||||
Definition sech x := 1 / cosh x.
|
||||
Definition csch x := 1 / sinh x.
|
||||
variable sin : Real → Real.
|
||||
definition cos x := sin (x - π / 2).
|
||||
definition tan x := sin x / cos x.
|
||||
definition cot x := cos x / sin x.
|
||||
definition sec x := 1 / cos x.
|
||||
definition csc x := 1 / sin x.
|
||||
definition sinh x := (1 - exp (-2 * x)) / (2 * exp (- x)).
|
||||
definition cosh x := (1 + exp (-2 * x)) / (2 * exp (- x)).
|
||||
definition tanh x := sinh x / cosh x.
|
||||
definition coth x := cosh x / sinh x.
|
||||
definition sech x := 1 / cosh x.
|
||||
definition csch x := 1 / sinh x.
|
||||
|
|
|
@ -17,7 +17,7 @@ class coercion_declaration : public neutral_object_cell {
|
|||
public:
|
||||
coercion_declaration(expr const & c):m_coercion(c) {}
|
||||
virtual ~coercion_declaration() {}
|
||||
virtual char const * keyword() const { return "Coercion"; }
|
||||
virtual char const * keyword() const { return "coercion"; }
|
||||
expr const & get_coercion() const { return m_coercion; }
|
||||
virtual void write(serializer & s) const;
|
||||
};
|
||||
|
|
|
@ -111,15 +111,15 @@ operator_info mixfixo(unsigned num_parts, name const * parts, unsigned precedenc
|
|||
|
||||
char const * to_string(fixity f) {
|
||||
switch (f) {
|
||||
case fixity::Infix: return "Infix";
|
||||
case fixity::Infixl: return "Infixl";
|
||||
case fixity::Infixr: return "Infixr";
|
||||
case fixity::Prefix: return "Prefix";
|
||||
case fixity::Postfix: return "Postfix";
|
||||
case fixity::Mixfixl: return "Mixfixl";
|
||||
case fixity::Mixfixr: return "Mixfixr";
|
||||
case fixity::Mixfixc: return "Mixfixc";
|
||||
case fixity::Mixfixo: return "Mixfixo";
|
||||
case fixity::Infix: return "infix";
|
||||
case fixity::Infixl: return "infixl";
|
||||
case fixity::Infixr: return "infixr";
|
||||
case fixity::Prefix: return "prefix";
|
||||
case fixity::Postfix: return "postfix";
|
||||
case fixity::Mixfixl: return "mixfixl";
|
||||
case fixity::Mixfixr: return "mixfixr";
|
||||
case fixity::Mixfixc: return "mixfixc";
|
||||
case fixity::Mixfixo: return "mixfixo";
|
||||
}
|
||||
lean_unreachable(); // LCOV_EXCL_LINE
|
||||
}
|
||||
|
@ -141,7 +141,7 @@ format pp(operator_info const & o) {
|
|||
case fixity::Mixfixr:
|
||||
case fixity::Mixfixc:
|
||||
case fixity::Mixfixo:
|
||||
r = highlight_command(format("Notation"));
|
||||
r = highlight_command(format("notation"));
|
||||
if (o.get_precedence() > 1)
|
||||
r += format{space(), format(o.get_precedence())};
|
||||
switch (o.get_fixity()) {
|
||||
|
@ -220,12 +220,12 @@ io_state_stream const & operator<<(io_state_stream const & out, operator_info co
|
|||
return out;
|
||||
}
|
||||
|
||||
char const * alias_declaration::keyword() const { return "Alias"; }
|
||||
void alias_declaration::write(serializer & s) const { s << "Alias" << m_name << m_expr; }
|
||||
char const * alias_declaration::keyword() const { return "alias"; }
|
||||
void alias_declaration::write(serializer & s) const { s << "alias" << m_name << m_expr; }
|
||||
static void read_alias(environment const & env, io_state const &, deserializer & d) {
|
||||
name n = read_name(d);
|
||||
expr e = read_expr(d);
|
||||
add_alias(env, n, e);
|
||||
}
|
||||
static object_cell::register_deserializer_fn add_alias_ds("Alias", read_alias);
|
||||
static object_cell::register_deserializer_fn add_alias_ds("alias", read_alias);
|
||||
}
|
||||
|
|
|
@ -24,41 +24,39 @@ Author: Leonardo de Moura
|
|||
namespace lean {
|
||||
// ==========================================
|
||||
// Builtin commands
|
||||
static name g_alias_kwd("Alias");
|
||||
static name g_definition_kwd("Definition");
|
||||
static name g_variable_kwd("Variable");
|
||||
static name g_variables_kwd("Variables");
|
||||
static name g_theorem_kwd("Theorem");
|
||||
static name g_axiom_kwd("Axiom");
|
||||
static name g_universe_kwd("Universe");
|
||||
static name g_eval_kwd("Eval");
|
||||
static name g_check_kwd("Check");
|
||||
static name g_infix_kwd("Infix");
|
||||
static name g_infixl_kwd("Infixl");
|
||||
static name g_infixr_kwd("Infixr");
|
||||
static name g_notation_kwd("Notation");
|
||||
static name g_set_option_kwd("SetOption");
|
||||
static name g_set_opaque_kwd("SetOpaque");
|
||||
static name g_options_kwd("Options");
|
||||
static name g_env_kwd("Environment");
|
||||
static name g_import_kwd("Import");
|
||||
static name g_alias_kwd("alias");
|
||||
static name g_definition_kwd("definition");
|
||||
static name g_variable_kwd("variable");
|
||||
static name g_variables_kwd("variables");
|
||||
static name g_theorem_kwd("theorem");
|
||||
static name g_axiom_kwd("axiom");
|
||||
static name g_universe_kwd("universe");
|
||||
static name g_eval_kwd("eval");
|
||||
static name g_check_kwd("check");
|
||||
static name g_infix_kwd("infix");
|
||||
static name g_infixl_kwd("infixl");
|
||||
static name g_infixr_kwd("infixr");
|
||||
static name g_notation_kwd("notation");
|
||||
static name g_set_option_kwd("setoption");
|
||||
static name g_set_opaque_kwd("setopaque");
|
||||
static name g_options_kwd("options");
|
||||
static name g_env_kwd("environment");
|
||||
static name g_import_kwd("import");
|
||||
static name g_help_kwd("help");
|
||||
static name g_coercion_kwd("Coercion");
|
||||
static name g_exit_kwd("Exit");
|
||||
static name g_coercion_kwd("coercion");
|
||||
static name g_exit_kwd("exit");
|
||||
static name g_print_kwd("print");
|
||||
static name g_push_kwd("Push");
|
||||
static name g_pop_kwd("Pop");
|
||||
static name g_scope_kwd("Scope");
|
||||
static name g_end_scope_kwd("EndScope");
|
||||
static name g_builtin_kwd("Builtin");
|
||||
static name g_namespace_kwd("Namespace");
|
||||
static name g_end_namespace_kwd("EndNamespace");
|
||||
static name g_pop_kwd("pop", "scope");
|
||||
static name g_scope_kwd("scope");
|
||||
static name g_builtin_kwd("builtin");
|
||||
static name g_namespace_kwd("namespace");
|
||||
static name g_end_kwd("end");
|
||||
/** \brief Table/List with all builtin command keywords */
|
||||
static list<name> g_command_keywords = {g_definition_kwd, g_variable_kwd, g_variables_kwd, g_theorem_kwd, g_axiom_kwd, g_universe_kwd, g_eval_kwd,
|
||||
g_check_kwd, g_infix_kwd, g_infixl_kwd, g_infixr_kwd, g_notation_kwd,
|
||||
g_set_option_kwd, g_set_opaque_kwd, g_env_kwd, g_options_kwd, g_import_kwd, g_help_kwd, g_coercion_kwd,
|
||||
g_exit_kwd, g_print_kwd, g_push_kwd, g_pop_kwd, g_scope_kwd, g_end_scope_kwd, g_alias_kwd, g_builtin_kwd,
|
||||
g_namespace_kwd, g_end_namespace_kwd};
|
||||
g_exit_kwd, g_print_kwd, g_pop_kwd, g_scope_kwd, g_alias_kwd, g_builtin_kwd,
|
||||
g_namespace_kwd, g_end_kwd};
|
||||
// ==========================================
|
||||
|
||||
list<name> const & parser_imp::get_command_keywords() {
|
||||
|
@ -327,7 +325,7 @@ void parser_imp::parse_print() {
|
|||
std::reverse(to_display.begin(), to_display.end());
|
||||
for (auto obj : to_display) {
|
||||
if (is_begin_import(obj)) {
|
||||
regular(m_io_state) << "Import \"" << *get_imported_module(obj) << "\"" << endl;
|
||||
regular(m_io_state) << "import \"" << *get_imported_module(obj) << "\"" << endl;
|
||||
} else {
|
||||
regular(m_io_state) << obj << endl;
|
||||
}
|
||||
|
@ -598,30 +596,28 @@ void parser_imp::parse_help() {
|
|||
}
|
||||
} else {
|
||||
regular(m_io_state) << "Available commands:" << endl
|
||||
<< " Alias [id] : [expr] define an alias for the given expression" << endl
|
||||
<< " Axiom [id] : [type] assert/postulate a new axiom" << endl
|
||||
<< " Check [expr] type check the given expression" << endl
|
||||
<< " Definition [id] : [type] := [expr] define a new element" << endl
|
||||
<< " Echo [string] display the given string" << endl
|
||||
<< " EndScope end the current scope and import its objects into the parent scope" << endl
|
||||
<< " Eval [expr] evaluate the given expression" << endl
|
||||
<< " Exit exit" << endl
|
||||
<< " Help display this message" << endl
|
||||
<< " Help Options display available options" << endl
|
||||
<< " Help Notation describe commands for defining infix, mixfix, postfix operators" << endl
|
||||
<< " Import [string] load the given file" << endl
|
||||
<< " Push create a scope (it is just an alias for the command Scope)" << endl
|
||||
<< " Pop discard the current scope" << endl
|
||||
<< " alias [id] : [expr] define an alias for the given expression" << endl
|
||||
<< " axiom [id] : [type] assert/postulate a new axiom" << endl
|
||||
<< " check [expr] type check the given expression" << endl
|
||||
<< " definition [id] : [type] := [expr] define a new element" << endl
|
||||
<< " end end the current scope/namespace" << endl
|
||||
<< " eval [expr] evaluate the given expression" << endl
|
||||
<< " exit exit" << endl
|
||||
<< " help display this message" << endl
|
||||
<< " help options display available options" << endl
|
||||
<< " help notation describe commands for defining infix, mixfix, postfix operators" << endl
|
||||
<< " import [string] load the given file" << endl
|
||||
<< " pop::scope discard the current scope" << endl
|
||||
<< " print [expr] pretty print the given expression" << endl
|
||||
<< " print Options print current the set of assigned options" << endl
|
||||
<< " print [string] print the given string" << endl
|
||||
<< " print Environment print objects in the environment, if [Num] provided, then show only the last [Num] objects" << endl
|
||||
<< " print Environment [num] show the last num objects in the environment" << endl
|
||||
<< " Scope create a scope" << endl
|
||||
<< " SetOption [id] [value] set option [id] with value [value]" << endl
|
||||
<< " Theorem [id] : [type] := [expr] define a new theorem" << endl
|
||||
<< " Variable [id] : [type] declare/postulate an element of the given type" << endl
|
||||
<< " Universe [id] [level] declare a new universe variable that is >= the given level" << endl;
|
||||
<< " scope create a scope" << endl
|
||||
<< " setoption [id] [value] set option [id] with value [value]" << endl
|
||||
<< " theorem [id] : [type] := [expr] define a new theorem" << endl
|
||||
<< " variable [id] : [type] declare/postulate an element of the given type" << endl
|
||||
<< " universe [id] [level] declare a new universe variable that is >= the given level" << endl;
|
||||
#if !defined(LEAN_WINDOWS)
|
||||
regular(m_io_state) << "Type Ctrl-D to exit" << endl;
|
||||
#endif
|
||||
|
@ -643,32 +639,6 @@ void parser_imp::reset_env(environment env) {
|
|||
m_io_state.set_formatter(mk_pp_formatter(env));
|
||||
}
|
||||
|
||||
void parser_imp::parse_scope() {
|
||||
next();
|
||||
reset_env(m_env->mk_child());
|
||||
}
|
||||
|
||||
void parser_imp::parse_pop() {
|
||||
next();
|
||||
if (!m_env->has_parent())
|
||||
throw parser_error("main scope cannot be removed", m_last_cmd_pos);
|
||||
reset_env(m_env->parent());
|
||||
}
|
||||
|
||||
void parser_imp::parse_end_scope() {
|
||||
next();
|
||||
if (!m_env->has_parent())
|
||||
throw parser_error("main scope cannot be removed", m_last_cmd_pos);
|
||||
auto new_objects = export_local_objects(m_env);
|
||||
reset_env(m_env->parent());
|
||||
for (auto const & obj : new_objects) {
|
||||
if (obj.is_theorem())
|
||||
m_env->add_theorem(obj.get_name(), obj.get_type(), obj.get_value());
|
||||
else
|
||||
m_env->add_definition(obj.get_name(), obj.get_type(), obj.get_value(), obj.is_opaque());
|
||||
}
|
||||
}
|
||||
|
||||
void parser_imp::parse_cmd_macro(name cmd_id, pos_info const & p) {
|
||||
lean_assert(m_cmd_macros && m_cmd_macros->find(cmd_id) != m_cmd_macros->end());
|
||||
next();
|
||||
|
@ -732,17 +702,56 @@ void parser_imp::parse_builtin() {
|
|||
register_implicit_arguments(full_id, parameters);
|
||||
}
|
||||
|
||||
void parser_imp::parse_scope() {
|
||||
next();
|
||||
m_scope_kinds.push_back(scope_kind::Scope);
|
||||
reset_env(m_env->mk_child());
|
||||
}
|
||||
|
||||
void parser_imp::parse_pop() {
|
||||
next();
|
||||
if (m_scope_kinds.empty())
|
||||
throw parser_error("main scope cannot be removed", m_last_cmd_pos);
|
||||
if (m_scope_kinds.back() != scope_kind::Scope)
|
||||
throw parser_error("invalid pop command, it is not inside of a scope", m_last_cmd_pos);
|
||||
if (!m_env->has_parent())
|
||||
throw parser_error("main scope cannot be removed", m_last_cmd_pos);
|
||||
m_scope_kinds.pop_back();
|
||||
reset_env(m_env->parent());
|
||||
}
|
||||
|
||||
void parser_imp::parse_namespace() {
|
||||
next();
|
||||
name id = check_identifier_next("invalid namespace declaration, identifier expected");
|
||||
m_scope_kinds.push_back(scope_kind::Namespace);
|
||||
m_namespace_prefixes.push_back(m_namespace_prefixes.back() + id);
|
||||
}
|
||||
|
||||
void parser_imp::parse_end_namespace() {
|
||||
void parser_imp::parse_end() {
|
||||
next();
|
||||
if (m_namespace_prefixes.size() <= 1)
|
||||
throw parser_error("invalid end namespace command, there are no open namespaces", m_last_cmd_pos);
|
||||
m_namespace_prefixes.pop_back();
|
||||
if (m_scope_kinds.empty())
|
||||
throw parser_error("invalid 'end', not inside of a scope or namespace", m_last_cmd_pos);
|
||||
scope_kind k = m_scope_kinds.back();
|
||||
m_scope_kinds.pop_back();
|
||||
switch (k) {
|
||||
case scope_kind::Scope: {
|
||||
if (!m_env->has_parent())
|
||||
throw parser_error("main scope cannot be removed", m_last_cmd_pos);
|
||||
auto new_objects = export_local_objects(m_env);
|
||||
reset_env(m_env->parent());
|
||||
for (auto const & obj : new_objects) {
|
||||
if (obj.is_theorem())
|
||||
m_env->add_theorem(obj.get_name(), obj.get_type(), obj.get_value());
|
||||
else
|
||||
m_env->add_definition(obj.get_name(), obj.get_type(), obj.get_value(), obj.is_opaque());
|
||||
}
|
||||
break;
|
||||
}
|
||||
case scope_kind::Namespace:
|
||||
if (m_namespace_prefixes.size() <= 1)
|
||||
throw parser_error("invalid end namespace command, there are no open namespaces", m_last_cmd_pos);
|
||||
m_namespace_prefixes.pop_back();
|
||||
}
|
||||
}
|
||||
|
||||
/** \brief Parse a Lean command. */
|
||||
|
@ -789,12 +798,10 @@ bool parser_imp::parse_command() {
|
|||
} else if (cmd_id == g_exit_kwd) {
|
||||
next();
|
||||
return false;
|
||||
} else if (cmd_id == g_push_kwd || cmd_id == g_scope_kwd) {
|
||||
} else if (cmd_id == g_scope_kwd) {
|
||||
parse_scope();
|
||||
} else if (cmd_id == g_pop_kwd) {
|
||||
parse_pop();
|
||||
} else if (cmd_id == g_end_scope_kwd) {
|
||||
parse_end_scope();
|
||||
} else if (cmd_id == g_universe_kwd) {
|
||||
parse_universe();
|
||||
} else if (cmd_id == g_alias_kwd) {
|
||||
|
@ -803,8 +810,8 @@ bool parser_imp::parse_command() {
|
|||
parse_builtin();
|
||||
} else if (cmd_id == g_namespace_kwd) {
|
||||
parse_namespace();
|
||||
} else if (cmd_id == g_end_namespace_kwd) {
|
||||
parse_end_namespace();
|
||||
} else if (cmd_id == g_end_kwd) {
|
||||
parse_end();
|
||||
} else if (m_cmd_macros && m_cmd_macros->find(cmd_id) != m_cmd_macros->end()) {
|
||||
parse_cmd_macro(cmd_id, m_last_cmd_pos);
|
||||
} else {
|
||||
|
|
|
@ -73,6 +73,8 @@ class parser_imp {
|
|||
pos_info m_last_script_pos;
|
||||
tactic_hints m_tactic_hints;
|
||||
std::vector<name> m_namespace_prefixes;
|
||||
enum class scope_kind { Scope, Namespace };
|
||||
std::vector<scope_kind> m_scope_kinds;
|
||||
|
||||
std::unique_ptr<calc_proof_parser> m_calc_proof_parser;
|
||||
|
||||
|
@ -409,13 +411,12 @@ private:
|
|||
void reset_env(environment env);
|
||||
void parse_scope();
|
||||
void parse_pop();
|
||||
void parse_end_scope();
|
||||
void parse_cmd_macro(name cmd_id, pos_info const & p);
|
||||
void parse_universe();
|
||||
void parse_alias();
|
||||
void parse_builtin();
|
||||
void parse_namespace();
|
||||
void parse_end_namespace();
|
||||
void parse_end();
|
||||
bool parse_command();
|
||||
/*@}*/
|
||||
|
||||
|
|
|
@ -33,7 +33,7 @@ class set_opaque_command : public neutral_object_cell {
|
|||
public:
|
||||
set_opaque_command(name const & n, bool opaque):m_obj_name(n), m_opaque(opaque) {}
|
||||
virtual ~set_opaque_command() {}
|
||||
virtual char const * keyword() const { return "SetOpaque"; }
|
||||
virtual char const * keyword() const { return "setopaque"; }
|
||||
virtual void write(serializer & s) const { s << "Opa" << m_obj_name << m_opaque; }
|
||||
name const & get_obj_name() const { return m_obj_name; }
|
||||
bool get_flag() const { return m_opaque; }
|
||||
|
@ -64,15 +64,15 @@ class import_command : public neutral_object_cell {
|
|||
public:
|
||||
import_command(std::string const & n):m_mod_name(n) {}
|
||||
virtual ~import_command() {}
|
||||
virtual char const * keyword() const { return "Import"; }
|
||||
virtual void write(serializer & s) const { s << "Import" << m_mod_name; }
|
||||
virtual char const * keyword() const { return "import"; }
|
||||
virtual void write(serializer & s) const { s << "import" << m_mod_name; }
|
||||
std::string const & get_module() const { return m_mod_name; }
|
||||
};
|
||||
static void read_import(environment const & env, io_state const & ios, deserializer & d) {
|
||||
std::string n = d.read_string();
|
||||
env->import(n, ios);
|
||||
}
|
||||
static object_cell::register_deserializer_fn import_ds("Import", read_import);
|
||||
static object_cell::register_deserializer_fn import_ds("import", read_import);
|
||||
|
||||
class end_import_mark : public neutral_object_cell {
|
||||
public:
|
||||
|
|
|
@ -58,15 +58,15 @@ public:
|
|||
|
||||
virtual bool has_cnstr_level() const { return true; }
|
||||
virtual level get_cnstr_level() const { return m_level; }
|
||||
virtual char const * keyword() const { return "Universe"; }
|
||||
virtual void write(serializer & s) const { s << "Universe" << get_name() << m_level; }
|
||||
virtual char const * keyword() const { return "universe"; }
|
||||
virtual void write(serializer & s) const { s << "universe" << get_name() << m_level; }
|
||||
};
|
||||
static void read_uvar_decl(environment const & env, io_state const &, deserializer & d) {
|
||||
name n = read_name(d);
|
||||
level lvl = read_level(d);
|
||||
env->add_uvar(n, lvl);
|
||||
}
|
||||
static object_cell::register_deserializer_fn uvar_ds("Universe", read_uvar_decl);
|
||||
static object_cell::register_deserializer_fn uvar_ds("universe", read_uvar_decl);
|
||||
|
||||
/**
|
||||
\brief Builtin object.
|
||||
|
@ -86,15 +86,15 @@ public:
|
|||
virtual bool is_opaque() const { return m_opaque; }
|
||||
virtual void set_opaque(bool f) { m_opaque = f; }
|
||||
virtual expr get_value() const { return m_value; }
|
||||
virtual char const * keyword() const { return "Builtin"; }
|
||||
virtual char const * keyword() const { return "builtin"; }
|
||||
virtual bool is_builtin() const { return true; }
|
||||
virtual void write(serializer & s) const { s << "Builtin" << m_value; }
|
||||
virtual void write(serializer & s) const { s << "builtin" << m_value; }
|
||||
};
|
||||
static void read_builtin(environment const & env, io_state const &, deserializer & d) {
|
||||
expr v = read_expr(d);
|
||||
env->add_builtin(v);
|
||||
}
|
||||
static object_cell::register_deserializer_fn builtin_ds("Builtin", read_builtin);
|
||||
static object_cell::register_deserializer_fn builtin_ds("builtin", read_builtin);
|
||||
|
||||
|
||||
/**
|
||||
|
@ -118,13 +118,13 @@ public:
|
|||
virtual name get_name() const { return to_value(m_representative).get_name(); }
|
||||
virtual bool is_builtin_set() const { return true; }
|
||||
virtual bool in_builtin_set(expr const & v) const { return is_value(v) && typeid(to_value(v)) == typeid(to_value(m_representative)); }
|
||||
virtual char const * keyword() const { return "BuiltinSet"; }
|
||||
virtual void write(serializer & s) const { s << "BuiltinSet" << m_representative; }
|
||||
virtual char const * keyword() const { return "builtinset"; }
|
||||
virtual void write(serializer & s) const { s << "builtinset" << m_representative; }
|
||||
};
|
||||
static void read_builtin_set(environment const & env, io_state const &, deserializer & d) {
|
||||
env->add_builtin_set(read_expr(d));
|
||||
}
|
||||
static object_cell::register_deserializer_fn builtin_set_ds("BuiltinSet", read_builtin_set);
|
||||
static object_cell::register_deserializer_fn builtin_set_ds("builtinset", read_builtin_set);
|
||||
|
||||
/**
|
||||
\brief Named (and typed) kernel objects.
|
||||
|
@ -155,16 +155,16 @@ public:
|
|||
class axiom_object_cell : public postulate_object_cell {
|
||||
public:
|
||||
axiom_object_cell(name const & n, expr const & t):postulate_object_cell(n, t) {}
|
||||
virtual char const * keyword() const { return "Axiom"; }
|
||||
virtual char const * keyword() const { return "axiom"; }
|
||||
virtual bool is_axiom() const { return true; }
|
||||
virtual void write(serializer & s) const { s << "Ax" << get_name() << get_type(); }
|
||||
virtual void write(serializer & s) const { s << "ax" << get_name() << get_type(); }
|
||||
};
|
||||
static void read_axiom(environment const & env, io_state const &, deserializer & d) {
|
||||
name n = read_name(d);
|
||||
expr t = read_expr(d);
|
||||
env->add_axiom(n, t);
|
||||
}
|
||||
static object_cell::register_deserializer_fn axiom_ds("Ax", read_axiom);
|
||||
static object_cell::register_deserializer_fn axiom_ds("ax", read_axiom);
|
||||
|
||||
|
||||
/**
|
||||
|
@ -173,16 +173,16 @@ static object_cell::register_deserializer_fn axiom_ds("Ax", read_axiom);
|
|||
class variable_decl_object_cell : public postulate_object_cell {
|
||||
public:
|
||||
variable_decl_object_cell(name const & n, expr const & t):postulate_object_cell(n, t) {}
|
||||
virtual char const * keyword() const { return "Variable"; }
|
||||
virtual char const * keyword() const { return "variable"; }
|
||||
virtual bool is_var_decl() const { return true; }
|
||||
virtual void write(serializer & s) const { s << "Var" << get_name() << get_type(); }
|
||||
virtual void write(serializer & s) const { s << "var" << get_name() << get_type(); }
|
||||
};
|
||||
static void read_variable(environment const & env, io_state const &, deserializer & d) {
|
||||
name n = read_name(d);
|
||||
expr t = read_expr(d);
|
||||
env->add_var(n, t);
|
||||
}
|
||||
static object_cell::register_deserializer_fn var_decl_ds("Var", read_variable);
|
||||
static object_cell::register_deserializer_fn var_decl_ds("var", read_variable);
|
||||
|
||||
/**
|
||||
\brief Base class for definitions: theorems and definitions.
|
||||
|
@ -200,9 +200,9 @@ public:
|
|||
virtual bool is_opaque() const { return m_opaque; }
|
||||
virtual void set_opaque(bool f) { m_opaque = f; }
|
||||
virtual expr get_value() const { return m_value; }
|
||||
virtual char const * keyword() const { return "Definition"; }
|
||||
virtual char const * keyword() const { return "definition"; }
|
||||
virtual unsigned get_weight() const { return m_weight; }
|
||||
virtual void write(serializer & s) const { s << "Def" << get_name() << get_type() << get_value(); }
|
||||
virtual void write(serializer & s) const { s << "def" << get_name() << get_type() << get_value(); }
|
||||
};
|
||||
static void read_definition(environment const & env, io_state const &, deserializer & d) {
|
||||
name n = read_name(d);
|
||||
|
@ -210,7 +210,7 @@ static void read_definition(environment const & env, io_state const &, deseriali
|
|||
expr v = read_expr(d);
|
||||
env->add_definition(n, t, v);
|
||||
}
|
||||
static object_cell::register_deserializer_fn definition_ds("Def", read_definition);
|
||||
static object_cell::register_deserializer_fn definition_ds("def", read_definition);
|
||||
|
||||
/**
|
||||
\brief Theorems
|
||||
|
@ -221,9 +221,9 @@ public:
|
|||
definition_object_cell(n, t, v, 0) {
|
||||
set_opaque(true);
|
||||
}
|
||||
virtual char const * keyword() const { return "Theorem"; }
|
||||
virtual char const * keyword() const { return "theorem"; }
|
||||
virtual bool is_theorem() const { return true; }
|
||||
virtual void write(serializer & s) const { s << "Th" << get_name() << get_type() << get_value(); }
|
||||
virtual void write(serializer & s) const { s << "th" << get_name() << get_type() << get_value(); }
|
||||
};
|
||||
static void read_theorem(environment const & env, io_state const &, deserializer & d) {
|
||||
name n = read_name(d);
|
||||
|
@ -231,7 +231,7 @@ static void read_theorem(environment const & env, io_state const &, deserializer
|
|||
expr v = read_expr(d);
|
||||
env->add_theorem(n, t, v);
|
||||
}
|
||||
static object_cell::register_deserializer_fn theorem_ds("Th", read_theorem);
|
||||
static object_cell::register_deserializer_fn theorem_ds("th", read_theorem);
|
||||
|
||||
object mk_uvar_decl(name const & n, level const & l) { return object(new uvar_declaration_object_cell(n, l)); }
|
||||
object mk_definition(name const & n, expr const & t, expr const & v, unsigned weight) { return object(new definition_object_cell(n, t, v, weight)); }
|
||||
|
|
|
@ -44,13 +44,13 @@ static void parse_error(environment const & env, io_state const & ios, char cons
|
|||
|
||||
static void tst1() {
|
||||
environment env; io_state ios = init_test_frontend(env);
|
||||
parse(env, ios, "Variable x : Bool Variable y : Bool Axiom H : x && y || x => x");
|
||||
parse(env, ios, "Eval true && true");
|
||||
parse(env, ios, "print true && false Eval true && false");
|
||||
parse(env, ios, "Infixl 35 & : and print true & false & false Eval true & false");
|
||||
parse(env, ios, "Notation 100 if _ then _ fi : implies print if true then false fi");
|
||||
parse(env, ios, "variable x : Bool variable y : Bool axiom H : x && y || x => x");
|
||||
parse(env, ios, "eval true && true");
|
||||
parse(env, ios, "print true && false eval true && false");
|
||||
parse(env, ios, "infixl 35 & : and print true & false & false eval true & false");
|
||||
parse(env, ios, "notation 100 if _ then _ fi : implies print if true then false fi");
|
||||
parse(env, ios, "print Pi (A : Type), A -> A");
|
||||
parse(env, ios, "Check Pi (A : Type), A -> A");
|
||||
parse(env, ios, "check Pi (A : Type), A -> A");
|
||||
}
|
||||
|
||||
static void check(environment const & env, io_state & ios, char const * str, expr const & expected) {
|
||||
|
@ -83,27 +83,27 @@ static void tst2() {
|
|||
static void tst3() {
|
||||
environment env; io_state ios = init_test_frontend(env);
|
||||
parse(env, ios, "help");
|
||||
parse(env, ios, "help Options");
|
||||
parse(env, ios, "help options");
|
||||
parse_error(env, ios, "help print");
|
||||
check(env, ios, "10.3", mk_real_value(mpq(103, 10)));
|
||||
parse(env, ios, "Variable f : Real -> Real. Check f 10.3.");
|
||||
parse(env, ios, "Variable g : (Type 1) -> Type. Check g Type");
|
||||
parse_error(env, ios, "Check fun .");
|
||||
parse_error(env, ios, "Definition foo .");
|
||||
parse_error(env, ios, "Check a");
|
||||
parse_error(env, ios, "Check U");
|
||||
parse(env, ios, "Variable h : Real -> Real -> Real. Notation 10 [ _ ; _ ] : h. Check [ 10.3 ; 20.1 ].");
|
||||
parse_error(env, ios, "Variable h : Real -> Real -> Real. Notation 10 [ _ ; _ ] : h. Check [ 10.3 | 20.1 ].");
|
||||
parse_error(env, ios, "SetOption pp::indent true");
|
||||
parse(env, ios, "SetOption pp::indent 10");
|
||||
parse_error(env, ios, "SetOption pp::colors foo");
|
||||
parse_error(env, ios, "SetOption pp::colors \"foo\"");
|
||||
parse(env, ios, "SetOption pp::colors true");
|
||||
parse_error(env, ios, "Notation 10 : Int::add");
|
||||
parse_error(env, ios, "Notation 10 _ : Int::add");
|
||||
parse(env, ios, "Notation 10 _ ++ _ : Int::add. Eval 10 ++ 20.");
|
||||
parse(env, ios, "Notation 10 _ -+ : Int::neg. Eval 10 -+");
|
||||
parse(env, ios, "Notation 30 -+ _ : Int::neg. Eval -+ 10");
|
||||
parse(env, ios, "variable f : Real -> Real. check f 10.3.");
|
||||
parse(env, ios, "variable g : (Type 1) -> Type. check g Type");
|
||||
parse_error(env, ios, "check fun .");
|
||||
parse_error(env, ios, "definition foo .");
|
||||
parse_error(env, ios, "check a");
|
||||
parse_error(env, ios, "check U");
|
||||
parse(env, ios, "variable h : Real -> Real -> Real. notation 10 [ _ ; _ ] : h. check [ 10.3 ; 20.1 ].");
|
||||
parse_error(env, ios, "variable h : Real -> Real -> Real. notation 10 [ _ ; _ ] : h. check [ 10.3 | 20.1 ].");
|
||||
parse_error(env, ios, "setoption pp::indent true");
|
||||
parse(env, ios, "setoption pp::indent 10");
|
||||
parse_error(env, ios, "setoption pp::colors foo");
|
||||
parse_error(env, ios, "setoption pp::colors \"foo\"");
|
||||
parse(env, ios, "setoption pp::colors true");
|
||||
parse_error(env, ios, "notation 10 : Int::add");
|
||||
parse_error(env, ios, "notation 10 _ : Int::add");
|
||||
parse(env, ios, "notation 10 _ ++ _ : Int::add. eval 10 ++ 20.");
|
||||
parse(env, ios, "notation 10 _ -+ : Int::neg. eval 10 -+");
|
||||
parse(env, ios, "notation 30 -+ _ : Int::neg. eval -+ 10");
|
||||
parse_error(env, ios, "10 + 30");
|
||||
}
|
||||
|
||||
|
|
|
@ -24,7 +24,7 @@ static void tst1() {
|
|||
environment env;
|
||||
env->add_var("N", Type());
|
||||
formatter fmt = mk_simple_formatter();
|
||||
check(fmt(env), "Variable N : Type\n");
|
||||
check(fmt(env), "variable N : Type\n");
|
||||
expr f = Const("f");
|
||||
expr a = Const("a");
|
||||
expr x = Const("x");
|
||||
|
@ -32,7 +32,7 @@ static void tst1() {
|
|||
expr N = Const("N");
|
||||
expr F = Fun({x, Pi({x, N}, x >> x)}, Let({y, f(a)}, f(Eq(x, f(y, a)))));
|
||||
check(fmt(F), "fun x : (Pi x : N, (x#0 -> x#1)), (let y := f a in (f (x#1 == (f y#0 a))))");
|
||||
check(fmt(env->get_object("N")), "Variable N : Type");
|
||||
check(fmt(env->get_object("N")), "variable N : Type");
|
||||
context ctx;
|
||||
ctx = extend(ctx, "x", f(a));
|
||||
ctx = extend(ctx, "y", f(Var(0), N >> N));
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
Import Int.
|
||||
Axiom PlusComm(a b : Int) : a + b == b + a.
|
||||
Variable a : Int.
|
||||
Check (Abst (fun x : Int, (PlusComm a x))).
|
||||
SetOption pp::implicit true.
|
||||
Check (Abst (fun x : Int, (PlusComm a x))).
|
||||
import Int.
|
||||
axiom PlusComm(a b : Int) : a + b == b + a.
|
||||
variable a : Int.
|
||||
check (Abst (fun x : Int, (PlusComm a x))).
|
||||
setoption pp::implicit true.
|
||||
check (Abst (fun x : Int, (PlusComm a x))).
|
||||
|
|
|
@ -1,3 +1,3 @@
|
|||
Alias BB : Bool.
|
||||
Variable x : BB.
|
||||
print Environment 1.
|
||||
alias BB : Bool.
|
||||
variable x : BB.
|
||||
print environment 1.
|
|
@ -1,4 +1,4 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Assumed: x
|
||||
Variable x : BB
|
||||
variable x : BB
|
||||
|
|
|
@ -1,16 +1,16 @@
|
|||
Push
|
||||
Variable Natural : Type.
|
||||
Alias ℕℕ : Natural.
|
||||
Variable x : Natural.
|
||||
print Environment 1.
|
||||
SetOption pp::unicode false.
|
||||
print Environment 1.
|
||||
SetOption pp::unicode true.
|
||||
print Environment 1.
|
||||
Alias NN : Natural.
|
||||
print Environment 2.
|
||||
Alias ℕℕℕ : Natural.
|
||||
print Environment 3.
|
||||
SetOption pp::unicode false.
|
||||
print Environment 3.
|
||||
Pop
|
||||
scope
|
||||
variable Natural : Type.
|
||||
alias ℕℕ : Natural.
|
||||
variable x : Natural.
|
||||
print environment 1.
|
||||
setoption pp::unicode false.
|
||||
print environment 1.
|
||||
setoption pp::unicode true.
|
||||
print environment 1.
|
||||
alias NN : Natural.
|
||||
print environment 2.
|
||||
alias ℕℕℕ : Natural.
|
||||
print environment 3.
|
||||
setoption pp::unicode false.
|
||||
print environment 3.
|
||||
pop::scope
|
|
@ -2,17 +2,17 @@
|
|||
Set: pp::unicode
|
||||
Assumed: Natural
|
||||
Assumed: x
|
||||
Variable x : ℕℕ
|
||||
variable x : ℕℕ
|
||||
Set: pp::unicode
|
||||
Variable x : Natural
|
||||
variable x : Natural
|
||||
Set: pp::unicode
|
||||
Variable x : ℕℕ
|
||||
Variable x : NN
|
||||
Alias NN : Natural
|
||||
Variable x : ℕℕℕ
|
||||
Alias NN : Natural
|
||||
Alias ℕℕℕ : Natural
|
||||
variable x : ℕℕ
|
||||
variable x : NN
|
||||
alias NN : Natural
|
||||
variable x : ℕℕℕ
|
||||
alias NN : Natural
|
||||
alias ℕℕℕ : Natural
|
||||
Set: pp::unicode
|
||||
Variable x : NN
|
||||
Alias NN : Natural
|
||||
Alias ℕℕℕ : Natural
|
||||
variable x : NN
|
||||
alias NN : Natural
|
||||
alias ℕℕℕ : Natural
|
||||
|
|
|
@ -1,25 +1,25 @@
|
|||
Import tactic.
|
||||
Import Int.
|
||||
import tactic.
|
||||
import Int.
|
||||
|
||||
Variable f : Int -> Int -> Bool
|
||||
Variable P : Int -> Int -> Bool
|
||||
Axiom Ax1 (x y : Int) (H : P x y) : (f x y)
|
||||
Theorem T1 (a : Int) : (P a a) => (f a a).
|
||||
variable f : Int -> Int -> Bool
|
||||
variable P : Int -> Int -> Bool
|
||||
axiom Ax1 (x y : Int) (H : P x y) : (f x y)
|
||||
theorem T1 (a : Int) : (P a a) => (f a a).
|
||||
apply Discharge.
|
||||
apply Ax1.
|
||||
exact.
|
||||
done.
|
||||
Variable b : Int
|
||||
Axiom Ax2 (x : Int) : (f x b)
|
||||
variable b : Int
|
||||
axiom Ax2 (x : Int) : (f x b)
|
||||
(*
|
||||
simple_tac = Repeat(OrElse(imp_tac(), assumption_tac(), apply_tac("Ax2"), apply_tac("Ax1")))
|
||||
*)
|
||||
Theorem T2 (a : Int) : (P a a) => (f a a).
|
||||
theorem T2 (a : Int) : (P a a) => (f a a).
|
||||
simple_tac.
|
||||
done.
|
||||
|
||||
Theorem T3 (a : Int) : (P a a) => (f a a).
|
||||
theorem T3 (a : Int) : (P a a) => (f a a).
|
||||
Repeat (OrElse (apply Discharge) exact (apply Ax2) (apply Ax1)).
|
||||
done.
|
||||
|
||||
print Environment 2.
|
||||
print environment 2.
|
|
@ -10,5 +10,5 @@
|
|||
Assumed: Ax2
|
||||
Proved: T2
|
||||
Proved: T3
|
||||
Theorem T2 (a : ℤ) : P a a ⇒ f a a := Discharge (λ H : P a a, Ax1 a a H)
|
||||
Theorem T3 (a : ℤ) : P a a ⇒ f a a := Discharge (λ H : P a a, Ax1 a a H)
|
||||
theorem T2 (a : ℤ) : P a a ⇒ f a a := Discharge (λ H : P a a, Ax1 a a H)
|
||||
theorem T3 (a : ℤ) : P a a ⇒ f a a := Discharge (λ H : P a a, Ax1 a a H)
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
(* import("tactic.lua") *)
|
||||
Check @Discharge
|
||||
Theorem T (a b : Bool) : a => b => b => a.
|
||||
check @Discharge
|
||||
theorem T (a b : Bool) : a => b => b => a.
|
||||
apply Discharge.
|
||||
apply Discharge.
|
||||
apply Discharge.
|
||||
|
|
|
@ -1,18 +1,18 @@
|
|||
Import Int.
|
||||
Check 10 + 20
|
||||
Check 10
|
||||
Check 10 - 20
|
||||
Eval 10 - 20
|
||||
Eval 15 + 10 - 20
|
||||
Check 15 + 10 - 20
|
||||
Variable x : Int
|
||||
Variable n : Nat
|
||||
Variable m : Nat
|
||||
import Int.
|
||||
check 10 + 20
|
||||
check 10
|
||||
check 10 - 20
|
||||
eval 10 - 20
|
||||
eval 15 + 10 - 20
|
||||
check 15 + 10 - 20
|
||||
variable x : Int
|
||||
variable n : Nat
|
||||
variable m : Nat
|
||||
print n + m
|
||||
print n + x + m
|
||||
SetOption lean::pp::coercion true
|
||||
setoption lean::pp::coercion true
|
||||
print n + x + m + 10
|
||||
print x + n + m + 10
|
||||
print n + m + 10 + x
|
||||
SetOption lean::pp::notation false
|
||||
setoption lean::pp::notation false
|
||||
print n + m + 10 + x
|
||||
|
|
|
@ -1,13 +1,13 @@
|
|||
Import Int.
|
||||
Import Real.
|
||||
import Int.
|
||||
import Real.
|
||||
print 1/2
|
||||
Eval 4/6
|
||||
eval 4/6
|
||||
print 3 div 2
|
||||
Variable x : Real
|
||||
Variable i : Int
|
||||
Variable n : Nat
|
||||
variable x : Real
|
||||
variable i : Int
|
||||
variable n : Nat
|
||||
print x + i + 1 + n
|
||||
SetOption lean::pp::coercion true
|
||||
setoption lean::pp::coercion true
|
||||
print x + i + 1 + n
|
||||
print x * i + x
|
||||
print x - i + x - x >= 0
|
||||
|
|
|
@ -1,10 +1,10 @@
|
|||
Import Int.
|
||||
Eval 8 mod 3
|
||||
Eval 8 div 4
|
||||
Eval 7 div 3
|
||||
Eval 7 mod 3
|
||||
import Int.
|
||||
eval 8 mod 3
|
||||
eval 8 div 4
|
||||
eval 7 div 3
|
||||
eval 7 mod 3
|
||||
print -8 mod 3
|
||||
SetOption lean::pp::notation false
|
||||
setoption lean::pp::notation false
|
||||
print -8 mod 3
|
||||
Eval -8 mod 3
|
||||
Eval (-8 div 3)*3 + (-8 mod 3)
|
||||
eval -8 mod 3
|
||||
eval (-8 div 3)*3 + (-8 mod 3)
|
|
@ -1,8 +1,8 @@
|
|||
Import specialfn.
|
||||
Variable x : Real
|
||||
Eval sin(x)
|
||||
Eval cos(x)
|
||||
Eval tan(x)
|
||||
Eval cot(x)
|
||||
Eval sec(x)
|
||||
Eval csc(x)
|
||||
import specialfn.
|
||||
variable x : Real
|
||||
eval sin(x)
|
||||
eval cos(x)
|
||||
eval tan(x)
|
||||
eval cot(x)
|
||||
eval sec(x)
|
||||
eval csc(x)
|
|
@ -1,8 +1,8 @@
|
|||
Import specialfn.
|
||||
Variable x : Real
|
||||
Eval sinh(x)
|
||||
Eval cosh(x)
|
||||
Eval tanh(x)
|
||||
Eval coth(x)
|
||||
Eval sech(x)
|
||||
Eval csch(x)
|
||||
import specialfn.
|
||||
variable x : Real
|
||||
eval sinh(x)
|
||||
eval cosh(x)
|
||||
eval tanh(x)
|
||||
eval coth(x)
|
||||
eval sech(x)
|
||||
eval csch(x)
|
|
@ -1,13 +1,13 @@
|
|||
Import Int.
|
||||
SetOption pp::unicode false
|
||||
import Int.
|
||||
setoption pp::unicode false
|
||||
print 3 | 6
|
||||
Eval 3 | 6
|
||||
Eval 3 | 7
|
||||
Eval 2 | 6
|
||||
Eval 1 | 6
|
||||
Variable x : Int
|
||||
Eval x | 3
|
||||
Eval 3 | x
|
||||
Eval 6 | 3
|
||||
SetOption pp::notation false
|
||||
eval 3 | 6
|
||||
eval 3 | 7
|
||||
eval 2 | 6
|
||||
eval 1 | 6
|
||||
variable x : Int
|
||||
eval x | 3
|
||||
eval 3 | x
|
||||
eval 6 | 3
|
||||
setoption pp::notation false
|
||||
print 3 | x
|
|
@ -1,16 +1,16 @@
|
|||
Import Int.
|
||||
Eval | -2 |
|
||||
import Int.
|
||||
eval | -2 |
|
||||
|
||||
-- Unfortunately, we can't write |-2|, because |- is considered a single token.
|
||||
-- It is not wise to change that since the symbol |- can be used as the notation for
|
||||
-- entailment relation in Lean.
|
||||
Eval |3|
|
||||
Definition x : Int := -3
|
||||
Eval |x + 1|
|
||||
Eval |x + 1| > 0
|
||||
Variable y : Int
|
||||
Eval |x + y|
|
||||
eval |3|
|
||||
definition x : Int := -3
|
||||
eval |x + 1|
|
||||
eval |x + 1| > 0
|
||||
variable y : Int
|
||||
eval |x + y|
|
||||
print |x + y| > x
|
||||
SetOption pp::notation false
|
||||
setoption pp::notation false
|
||||
print |x + y| > x
|
||||
print |x + y| + |y + x| > x
|
|
@ -1,6 +1,6 @@
|
|||
Import Real.
|
||||
Eval 10.3
|
||||
Eval 0.3
|
||||
Eval 0.3 + 0.1
|
||||
Eval 0.2 + 0.7
|
||||
Eval 1/3 + 0.1
|
||||
import Real.
|
||||
eval 10.3
|
||||
eval 0.3
|
||||
eval 0.3 + 0.1
|
||||
eval 0.2 + 0.7
|
||||
eval 1/3 + 0.1
|
|
@ -1,4 +1,4 @@
|
|||
Import Int.
|
||||
import Int.
|
||||
print (Int -> Int) -> Int
|
||||
print Int -> Int -> Int
|
||||
print Int -> (Int -> Int)
|
||||
|
|
|
@ -1,4 +1,4 @@
|
|||
Import Int.
|
||||
Variable g : Pi A : Type, A -> A.
|
||||
Variables a b : Int
|
||||
Axiom H1 : g _ a > 0
|
||||
import Int.
|
||||
variable g : Pi A : Type, A -> A.
|
||||
variables a b : Int
|
||||
axiom H1 : g _ a > 0
|
||||
|
|
|
@ -1,8 +1,8 @@
|
|||
SetOption pp::implicit true.
|
||||
SetOption pp::colors false.
|
||||
Variable N : Type.
|
||||
setoption pp::implicit true.
|
||||
setoption pp::colors false.
|
||||
variable N : Type.
|
||||
|
||||
Definition T (a : N) (f : _ -> _) (H : f a == a) : f (f _) == f _ :=
|
||||
definition T (a : N) (f : _ -> _) (H : f a == a) : f (f _) == f _ :=
|
||||
SubstP (fun x : N, f (f a) == _) (Refl (f (f _))) H.
|
||||
|
||||
print Environment 1.
|
||||
print environment 1.
|
||||
|
|
|
@ -4,5 +4,5 @@
|
|||
Set: pp::colors
|
||||
Assumed: N
|
||||
Defined: T
|
||||
Definition T (a : N) (f : N → N) (H : f a == a) : f (f a) == f (f a) :=
|
||||
definition T (a : N) (f : N → N) (H : f a == a) : f (f a) == f (f a) :=
|
||||
@SubstP N (f a) a (λ x : N, f (f a) == f (f a)) (@Refl N (f (f a))) H
|
||||
|
|
|
@ -1,13 +1,13 @@
|
|||
Import Int.
|
||||
Variable list : Type -> Type
|
||||
Variable nil {A : Type} : list A
|
||||
Variable cons {A : Type} (head : A) (tail : list A) : list A
|
||||
Definition n1 : list Int := cons (nat_to_int 10) nil
|
||||
Definition n2 : list Nat := cons 10 nil
|
||||
Definition n3 : list Int := cons 10 nil
|
||||
Definition n4 : list Int := nil
|
||||
Definition n5 : _ := cons 10 nil
|
||||
import Int.
|
||||
variable list : Type -> Type
|
||||
variable nil {A : Type} : list A
|
||||
variable cons {A : Type} (head : A) (tail : list A) : list A
|
||||
definition n1 : list Int := cons (nat_to_int 10) nil
|
||||
definition n2 : list Nat := cons 10 nil
|
||||
definition n3 : list Int := cons 10 nil
|
||||
definition n4 : list Int := nil
|
||||
definition n5 : _ := cons 10 nil
|
||||
|
||||
SetOption pp::coercion true
|
||||
SetOption pp::implicit true
|
||||
print Environment 1.
|
||||
setoption pp::coercion true
|
||||
setoption pp::implicit true
|
||||
print environment 1.
|
|
@ -11,4 +11,4 @@
|
|||
Defined: n5
|
||||
Set: lean::pp::coercion
|
||||
Set: lean::pp::implicit
|
||||
Definition n5 : list ℕ := @cons ℕ 10 (@nil ℕ)
|
||||
definition n5 : list ℕ := @cons ℕ 10 (@nil ℕ)
|
||||
|
|
|
@ -1,9 +1,9 @@
|
|||
Import Int.
|
||||
Variable list : Type -> Type
|
||||
Variable nil {A : Type} : list A
|
||||
Variable cons {A : Type} (head : A) (tail : list A) : list A
|
||||
Variables a b : Int
|
||||
Variables n m : Nat
|
||||
Definition l1 : list Int := cons a (cons b (cons n nil))
|
||||
Definition l2 : list Int := cons a (cons n (cons b nil))
|
||||
Check cons a (cons b (cons n nil))
|
||||
import Int.
|
||||
variable list : Type -> Type
|
||||
variable nil {A : Type} : list A
|
||||
variable cons {A : Type} (head : A) (tail : list A) : list A
|
||||
variables a b : Int
|
||||
variables n m : Nat
|
||||
definition l1 : list Int := cons a (cons b (cons n nil))
|
||||
definition l2 : list Int := cons a (cons n (cons b nil))
|
||||
check cons a (cons b (cons n nil))
|
||||
|
|
|
@ -1,4 +1,4 @@
|
|||
Import Int.
|
||||
Variable f {A : Type} (a : A) : A
|
||||
Variable a : Int
|
||||
Definition tst : Bool := (fun x, (f x) > 10) a
|
||||
import Int.
|
||||
variable f {A : Type} (a : A) : A
|
||||
variable a : Int
|
||||
definition tst : Bool := (fun x, (f x) > 10) a
|
||||
|
|
|
@ -1,8 +1,8 @@
|
|||
Import Int.
|
||||
Variable g {A : Type} (a : A) : A
|
||||
Variable a : Int
|
||||
Variable b : Int
|
||||
Axiom H1 : a = b
|
||||
Axiom H2 : (g a) > 0
|
||||
Theorem T1 : (g b) > 0 := SubstP (λ x, (g x) > 0) H2 H1
|
||||
print Environment 2
|
||||
import Int.
|
||||
variable g {A : Type} (a : A) : A
|
||||
variable a : Int
|
||||
variable b : Int
|
||||
axiom H1 : a = b
|
||||
axiom H2 : (g a) > 0
|
||||
theorem T1 : (g b) > 0 := SubstP (λ x, (g x) > 0) H2 H1
|
||||
print environment 2
|
||||
|
|
|
@ -7,5 +7,5 @@
|
|||
Assumed: H1
|
||||
Assumed: H2
|
||||
Proved: T1
|
||||
Axiom H2 : g a > 0
|
||||
Theorem T1 : g b > 0 := SubstP (λ x : ℤ, g x > 0) H2 H1
|
||||
axiom H2 : g a > 0
|
||||
theorem T1 : g b > 0 := SubstP (λ x : ℤ, g x > 0) H2 H1
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
Import Int.
|
||||
Import Real.
|
||||
Variable f {A : Type} (a : A) : A
|
||||
Variable a : Int
|
||||
Variable b : Real
|
||||
Definition tst : Bool := (fun x y, (f x) > (f y)) a b
|
||||
import Int.
|
||||
import Real.
|
||||
variable f {A : Type} (a : A) : A
|
||||
variable a : Int
|
||||
variable b : Real
|
||||
definition tst : Bool := (fun x y, (f x) > (f y)) a b
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
Import Real.
|
||||
Variable f {A : Type} (a b : A) : Bool
|
||||
Variable a : Int
|
||||
Variable b : Real
|
||||
Definition tst : Bool := (fun x y, f x y) a b
|
||||
print Environment 1
|
||||
import Real.
|
||||
variable f {A : Type} (a b : A) : Bool
|
||||
variable a : Int
|
||||
variable b : Real
|
||||
definition tst : Bool := (fun x y, f x y) a b
|
||||
print environment 1
|
||||
|
|
|
@ -5,4 +5,4 @@
|
|||
Assumed: a
|
||||
Assumed: b
|
||||
Defined: tst
|
||||
Definition tst : Bool := (λ x y : ℝ, f x y) a b
|
||||
definition tst : Bool := (λ x y : ℝ, f x y) a b
|
||||
|
|
|
@ -1,10 +1,10 @@
|
|||
Import Int.
|
||||
Variable list : Type → Type
|
||||
Variable nil {A : Type} : list A
|
||||
Variable cons {A : Type} (head : A) (tail : list A) : list A
|
||||
Variable a : ℤ
|
||||
Variable b : ℤ
|
||||
Variable n : ℕ
|
||||
Variable m : ℕ
|
||||
Definition l1 : list ℤ := cons a (cons b (cons n nil))
|
||||
Definition l2 : list ℤ := cons a (cons n (cons b nil))
|
||||
import Int.
|
||||
variable list : Type → Type
|
||||
variable nil {A : Type} : list A
|
||||
variable cons {A : Type} (head : A) (tail : list A) : list A
|
||||
variable a : ℤ
|
||||
variable b : ℤ
|
||||
variable n : ℕ
|
||||
variable m : ℕ
|
||||
definition l1 : list ℤ := cons a (cons b (cons n nil))
|
||||
definition l2 : list ℤ := cons a (cons n (cons b nil))
|
||||
|
|
|
@ -1,8 +1,8 @@
|
|||
SetOption pp::implicit true.
|
||||
SetOption pp::colors false.
|
||||
Variable N : Type.
|
||||
setoption pp::implicit true.
|
||||
setoption pp::colors false.
|
||||
variable N : Type.
|
||||
|
||||
Check
|
||||
check
|
||||
fun (a : N) (f : N -> N) (H : f a == a),
|
||||
let calc1 : f a == a := SubstP (fun x : N, f a == _) (Refl (f a)) H
|
||||
in calc1.
|
|
@ -1 +1 @@
|
|||
Check fun (A A' : (Type U)) (H : A == A'), Symm H
|
||||
check fun (A A' : (Type U)) (H : A == A'), Symm H
|
||||
|
|
|
@ -1,12 +1,12 @@
|
|||
Import tactic.
|
||||
import tactic.
|
||||
|
||||
Variables a b c d e : Bool.
|
||||
Axiom H1 : a => b.
|
||||
Axiom H2 : b = c.
|
||||
Axiom H3 : c => d.
|
||||
Axiom H4 : d <=> e.
|
||||
variables a b c d e : Bool.
|
||||
axiom H1 : a => b.
|
||||
axiom H2 : b = c.
|
||||
axiom H3 : c => d.
|
||||
axiom H4 : d <=> e.
|
||||
|
||||
Theorem T : a => e
|
||||
theorem T : a => e
|
||||
:= calc a => b : H1
|
||||
... = c : H2
|
||||
... => d : (by apply H3)
|
||||
|
|
|
@ -1,10 +1,10 @@
|
|||
|
||||
|
||||
Variables a b c d e : Nat.
|
||||
Variable f : Nat -> Nat.
|
||||
Axiom H1 : f a = a.
|
||||
variables a b c d e : Nat.
|
||||
variable f : Nat -> Nat.
|
||||
axiom H1 : f a = a.
|
||||
|
||||
Theorem T : f (f (f a)) = a
|
||||
theorem T : f (f (f a)) = a
|
||||
:= calc f (f (f a)) = f (f a) : { H1 }
|
||||
... = f a : { H1 }
|
||||
... = a : { H1 }.
|
||||
|
|
|
@ -1,19 +1,19 @@
|
|||
Import cast.
|
||||
Import Int.
|
||||
import cast.
|
||||
import Int.
|
||||
|
||||
Variable vector : Type -> Nat -> Type
|
||||
Axiom N0 (n : Nat) : n + 0 = n
|
||||
Theorem V0 (T : Type) (n : Nat) : (vector T (n + 0)) = (vector T n) :=
|
||||
variable vector : Type -> Nat -> Type
|
||||
axiom N0 (n : Nat) : n + 0 = n
|
||||
theorem V0 (T : Type) (n : Nat) : (vector T (n + 0)) = (vector T n) :=
|
||||
Congr (Refl (vector T)) (N0 n)
|
||||
Variable f (n : Nat) (v : vector Int n) : Int
|
||||
Variable m : Nat
|
||||
Variable v1 : vector Int (m + 0)
|
||||
variable f (n : Nat) (v : vector Int n) : Int
|
||||
variable m : Nat
|
||||
variable v1 : vector Int (m + 0)
|
||||
-- The following application will fail because (vector Int (m + 0)) and (vector Int m)
|
||||
-- are not definitionally equal.
|
||||
Check f m v1
|
||||
check f m v1
|
||||
-- The next one succeeds using the "casting" operator.
|
||||
-- We can do it, because (V0 Int m) is a proof that
|
||||
-- (vector Int (m + 0)) and (vector Int m) are propositionally equal.
|
||||
-- That is, they have the same interpretation in the lean set theoretic
|
||||
-- semantics.
|
||||
Check f m (cast (V0 Int m) v1)
|
||||
check f m (cast (V0 Int m) v1)
|
||||
|
|
|
@ -1,12 +1,12 @@
|
|||
Import cast
|
||||
Variable A : Type
|
||||
Variable B : Type
|
||||
Variable A' : Type
|
||||
Variable B' : Type
|
||||
Axiom H : (A -> B) = (A' -> B')
|
||||
Variable a : A
|
||||
Check DomInj H
|
||||
Theorem BeqB' : B = B' := RanInj H a
|
||||
SetOption pp::implicit true
|
||||
import cast
|
||||
variable A : Type
|
||||
variable B : Type
|
||||
variable A' : Type
|
||||
variable B' : Type
|
||||
axiom H : (A -> B) = (A' -> B')
|
||||
variable a : A
|
||||
check DomInj H
|
||||
theorem BeqB' : B = B' := RanInj H a
|
||||
setoption pp::implicit true
|
||||
print DomInj H
|
||||
print RanInj H a
|
||||
|
|
|
@ -1,24 +1,24 @@
|
|||
Import cast
|
||||
import cast
|
||||
|
||||
Variables A A' B B' : Type
|
||||
Variable x : A
|
||||
Eval cast (Refl A) x
|
||||
Eval x = (cast (Refl A) x)
|
||||
Variable b : B
|
||||
Definition f (x : A) : B := b
|
||||
Axiom H : (A -> B) = (A' -> B)
|
||||
Variable a' : A'
|
||||
Eval (cast H f) a'
|
||||
Axiom H2 : (A -> B) = (A' -> B')
|
||||
Definition g (x : B') : Nat := 0
|
||||
Eval g ((cast H2 f) a')
|
||||
Check g ((cast H2 f) a')
|
||||
variables A A' B B' : Type
|
||||
variable x : A
|
||||
eval cast (Refl A) x
|
||||
eval x = (cast (Refl A) x)
|
||||
variable b : B
|
||||
definition f (x : A) : B := b
|
||||
axiom H : (A -> B) = (A' -> B)
|
||||
variable a' : A'
|
||||
eval (cast H f) a'
|
||||
axiom H2 : (A -> B) = (A' -> B')
|
||||
definition g (x : B') : Nat := 0
|
||||
eval g ((cast H2 f) a')
|
||||
check g ((cast H2 f) a')
|
||||
|
||||
Eval (cast H2 f) a'
|
||||
eval (cast H2 f) a'
|
||||
|
||||
Variables A1 A2 A3 : Type
|
||||
Axiom Ha : A1 = A2
|
||||
Axiom Hb : A2 = A3
|
||||
Variable a : A1
|
||||
Eval (cast Hb (cast Ha a))
|
||||
Check (cast Hb (cast Ha a))
|
||||
variables A1 A2 A3 : Type
|
||||
axiom Ha : A1 = A2
|
||||
axiom Hb : A2 = A3
|
||||
variable a : A1
|
||||
eval (cast Hb (cast Ha a))
|
||||
check (cast Hb (cast Ha a))
|
||||
|
|
|
@ -1,7 +1,7 @@
|
|||
Import cast
|
||||
SetOption pp::colors false
|
||||
import cast
|
||||
setoption pp::colors false
|
||||
|
||||
Check fun (A A': TypeM)
|
||||
check fun (A A': TypeM)
|
||||
(B : A -> TypeM)
|
||||
(B' : A' -> TypeM)
|
||||
(f : Pi x : A, B x)
|
||||
|
|
|
@ -1,15 +1,15 @@
|
|||
Variable T : Type
|
||||
Variable R : Type
|
||||
Variable f : T -> R
|
||||
Coercion f
|
||||
print Environment 2
|
||||
Variable g : T -> R
|
||||
Coercion g
|
||||
Variable h : Pi (x : Type), x
|
||||
Coercion h
|
||||
Definition T2 : Type := T
|
||||
Definition R2 : Type := R
|
||||
Variable f2 : T2 -> R2
|
||||
Coercion f2
|
||||
Variable id : T -> T
|
||||
Coercion id
|
||||
variable T : Type
|
||||
variable R : Type
|
||||
variable f : T -> R
|
||||
coercion f
|
||||
print environment 2
|
||||
variable g : T -> R
|
||||
coercion g
|
||||
variable h : Pi (x : Type), x
|
||||
coercion h
|
||||
definition T2 : Type := T
|
||||
definition R2 : Type := R
|
||||
variable f2 : T2 -> R2
|
||||
coercion f2
|
||||
variable id : T -> T
|
||||
coercion id
|
||||
|
|
|
@ -4,8 +4,8 @@
|
|||
Assumed: R
|
||||
Assumed: f
|
||||
Coercion f
|
||||
Variable f : T → R
|
||||
Coercion f
|
||||
variable f : T → R
|
||||
coercion f
|
||||
Assumed: g
|
||||
Error (line: 8, pos: 0) invalid coercion declaration, frontend already has a coercion for the given types
|
||||
Assumed: h
|
||||
|
|
|
@ -1,32 +1,32 @@
|
|||
Variable T : Type
|
||||
Variable R : Type
|
||||
Variable t2r : T -> R
|
||||
Coercion t2r
|
||||
Variable g : R -> R -> R
|
||||
Variable a : T
|
||||
variable T : Type
|
||||
variable R : Type
|
||||
variable t2r : T -> R
|
||||
coercion t2r
|
||||
variable g : R -> R -> R
|
||||
variable a : T
|
||||
print g a a
|
||||
Variable b : R
|
||||
variable b : R
|
||||
print g a b
|
||||
print g b a
|
||||
SetOption lean::pp::coercion true
|
||||
setoption lean::pp::coercion true
|
||||
print g a a
|
||||
print g a b
|
||||
print g b a
|
||||
SetOption lean::pp::coercion false
|
||||
Variable S : Type
|
||||
Variable s : S
|
||||
Variable r : S
|
||||
Variable h : S -> S -> S
|
||||
Infixl 10 ++ : g
|
||||
Infixl 10 ++ : h
|
||||
SetOption lean::pp::notation false
|
||||
setoption lean::pp::coercion false
|
||||
variable S : Type
|
||||
variable s : S
|
||||
variable r : S
|
||||
variable h : S -> S -> S
|
||||
infixl 10 ++ : g
|
||||
infixl 10 ++ : h
|
||||
setoption lean::pp::notation false
|
||||
print a ++ b ++ a
|
||||
print r ++ s ++ r
|
||||
Check a ++ b ++ a
|
||||
Check r ++ s ++ r
|
||||
SetOption lean::pp::coercion true
|
||||
check a ++ b ++ a
|
||||
check r ++ s ++ r
|
||||
setoption lean::pp::coercion true
|
||||
print a ++ b ++ a
|
||||
print r ++ s ++ r
|
||||
SetOption lean::pp::notation true
|
||||
setoption lean::pp::notation true
|
||||
print a ++ b ++ a
|
||||
print r ++ s ++ r
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
Import specialfn.
|
||||
Definition f x y := x + y
|
||||
Definition g x y := sin x + y
|
||||
Definition h x y := x * sin (x + y)
|
||||
print Environment 3
|
||||
import specialfn.
|
||||
definition f x y := x + y
|
||||
definition g x y := sin x + y
|
||||
definition h x y := x * sin (x + y)
|
||||
print environment 3
|
|
@ -4,6 +4,6 @@
|
|||
Defined: f
|
||||
Defined: g
|
||||
Defined: h
|
||||
Definition f (x y : ℕ) : ℕ := x + y
|
||||
Definition g (x y : ℝ) : ℝ := sin x + y
|
||||
Definition h (x y : ℝ) : ℝ := x * sin (x + y)
|
||||
definition f (x y : ℕ) : ℕ := x + y
|
||||
definition g (x y : ℝ) : ℝ := sin x + y
|
||||
definition h (x y : ℝ) : ℝ := x * sin (x + y)
|
||||
|
|
|
@ -1,3 +1,3 @@
|
|||
-- SetOption default configuration for tests
|
||||
SetOption pp::colors false
|
||||
SetOption pp::unicode true
|
||||
-- setoption default configuration for tests
|
||||
setoption pp::colors false
|
||||
setoption pp::unicode true
|
||||
|
|
|
@ -1,20 +1,20 @@
|
|||
Import Int.
|
||||
Definition id (A : Type) : (Type U) := A.
|
||||
Variable p : (Int -> Int) -> Bool.
|
||||
Check fun (x : id Int), x.
|
||||
Variable f : (id Int) -> (id Int).
|
||||
Check p f.
|
||||
import Int.
|
||||
definition id (A : Type) : (Type U) := A.
|
||||
variable p : (Int -> Int) -> Bool.
|
||||
check fun (x : id Int), x.
|
||||
variable f : (id Int) -> (id Int).
|
||||
check p f.
|
||||
|
||||
Definition c (A : (Type 3)) : (Type 3) := (Type 1).
|
||||
Variable g : (c (Type 2)) -> Bool.
|
||||
Variable a : (c (Type 1)).
|
||||
Check g a.
|
||||
definition c (A : (Type 3)) : (Type 3) := (Type 1).
|
||||
variable g : (c (Type 2)) -> Bool.
|
||||
variable a : (c (Type 1)).
|
||||
check g a.
|
||||
|
||||
Definition c2 {T : Type} (A : (Type 3)) (a : T) : (Type 3) := (Type 1)
|
||||
Variable b : Int
|
||||
Check @c2
|
||||
Variable g2 : (c2 (Type 2) b) -> Bool
|
||||
Check g2
|
||||
Variable a2 : (c2 (Type 1) b).
|
||||
Check g2 a2
|
||||
Check fun x : (c2 (Type 1) b), g2 x
|
||||
definition c2 {T : Type} (A : (Type 3)) (a : T) : (Type 3) := (Type 1)
|
||||
variable b : Int
|
||||
check @c2
|
||||
variable g2 : (c2 (Type 2) b) -> Bool
|
||||
check g2
|
||||
variable a2 : (c2 (Type 1) b).
|
||||
check g2 a2
|
||||
check fun x : (c2 (Type 1) b), g2 x
|
|
@ -1,6 +1,6 @@
|
|||
Import tactic
|
||||
Check @Discharge
|
||||
Theorem T (a b : Bool) : a => b => b => a.
|
||||
import tactic
|
||||
check @Discharge
|
||||
theorem T (a b : Bool) : a => b => b => a.
|
||||
apply Discharge.
|
||||
apply Discharge.
|
||||
apply Discharge.
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
Import tactic
|
||||
import tactic
|
||||
|
||||
Theorem T1 (a b : Bool) : a \/ b => b \/ a.
|
||||
theorem T1 (a b : Bool) : a \/ b => b \/ a.
|
||||
apply Discharge.
|
||||
(* disj_hyp_tac() *)
|
||||
(* disj_tac() *)
|
||||
|
@ -14,8 +14,8 @@ Theorem T1 (a b : Bool) : a \/ b => b \/ a.
|
|||
simple_tac = Repeat(OrElse(imp_tac(), assumption_tac(), disj_hyp_tac(), disj_tac())) .. now_tac()
|
||||
*)
|
||||
|
||||
Theorem T2 (a b : Bool) : a \/ b => b \/ a.
|
||||
theorem T2 (a b : Bool) : a \/ b => b \/ a.
|
||||
simple_tac.
|
||||
done.
|
||||
|
||||
print Environment 1.
|
||||
print environment 1.
|
|
@ -3,5 +3,5 @@
|
|||
Imported 'tactic'
|
||||
Proved: T1
|
||||
Proved: T2
|
||||
Theorem T2 (a b : Bool) : a ∨ b ⇒ b ∨ a :=
|
||||
theorem T2 (a b : Bool) : a ∨ b ⇒ b ∨ a :=
|
||||
Discharge (λ H : a ∨ b, DisjCases H (λ H : a, Disj2 b H) (λ H : b, Disj1 H a))
|
||||
|
|
|
@ -1,7 +1,7 @@
|
|||
(* import("tactic.lua") *)
|
||||
Variables a b c : Bool
|
||||
Axiom H : a \/ b
|
||||
Theorem T (a b : Bool) : a \/ b => b \/ a.
|
||||
variables a b c : Bool
|
||||
axiom H : a \/ b
|
||||
theorem T (a b : Bool) : a \/ b => b \/ a.
|
||||
apply Discharge.
|
||||
apply (DisjCases H).
|
||||
apply Disj2.
|
||||
|
|
|
@ -1,29 +1,29 @@
|
|||
|
||||
|
||||
Variable f {A : Type} (a b : A) : A.
|
||||
Check f 10 true
|
||||
variable f {A : Type} (a b : A) : A.
|
||||
check f 10 true
|
||||
|
||||
Variable g {A B : Type} (a : A) : A.
|
||||
Check g 10
|
||||
variable g {A B : Type} (a : A) : A.
|
||||
check g 10
|
||||
|
||||
Variable h : Pi (A : Type), A -> A.
|
||||
variable h : Pi (A : Type), A -> A.
|
||||
|
||||
Check fun x, fun A : Type, h A x
|
||||
check fun x, fun A : Type, h A x
|
||||
|
||||
Variable my_eq : Pi A : Type, A -> A -> Bool.
|
||||
variable my_eq : Pi A : Type, A -> A -> Bool.
|
||||
|
||||
Check fun (A B : Type) (a : _) (b : _) (C : Type), my_eq C a b.
|
||||
check fun (A B : Type) (a : _) (b : _) (C : Type), my_eq C a b.
|
||||
|
||||
Variable a : Bool
|
||||
Variable b : Bool
|
||||
Variable H : a /\ b
|
||||
Theorem t1 : b := Discharge (fun H1, Conj H1 (Conjunct1 H)).
|
||||
variable a : Bool
|
||||
variable b : Bool
|
||||
variable H : a /\ b
|
||||
theorem t1 : b := Discharge (fun H1, Conj H1 (Conjunct1 H)).
|
||||
|
||||
Theorem t2 : a = b := Trans (Refl a) (Refl b).
|
||||
theorem t2 : a = b := Trans (Refl a) (Refl b).
|
||||
|
||||
Check f Bool Bool.
|
||||
check f Bool Bool.
|
||||
|
||||
Theorem pierce (a b : Bool) : ((a ⇒ b) ⇒ a) ⇒ a :=
|
||||
theorem pierce (a b : Bool) : ((a ⇒ b) ⇒ a) ⇒ a :=
|
||||
Discharge (λ H, DisjCases (EM a)
|
||||
(λ H_a, H)
|
||||
(λ H_na, NotImp1 (MT H H_na)))
|
||||
|
|
|
@ -1,24 +1,24 @@
|
|||
Variable C : Pi (A B : Type) (H : A = B) (a : A), B
|
||||
variable C : Pi (A B : Type) (H : A = B) (a : A), B
|
||||
|
||||
Variable D : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)), A = A'
|
||||
variable D : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)), A = A'
|
||||
|
||||
Variable R : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A),
|
||||
variable R : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A),
|
||||
(B a) = (B' (C A A' (D A A' B B' H) a))
|
||||
|
||||
Theorem R2 (A A' B B' : Type) (H : (A -> B) = (A' -> B')) (a : A) : B = B' := R _ _ _ _ H a
|
||||
theorem R2 (A A' B B' : Type) (H : (A -> B) = (A' -> B')) (a : A) : B = B' := R _ _ _ _ H a
|
||||
|
||||
print Environment 1
|
||||
print environment 1
|
||||
|
||||
Theorem R3 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
theorem R3 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
fun (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1),
|
||||
R _ _ _ _ H a
|
||||
|
||||
Theorem R4 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
theorem R4 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
fun (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : _),
|
||||
R _ _ _ _ H a
|
||||
|
||||
Theorem R5 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
theorem R5 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
fun (A1 A2 B1 B2 : Type) (H : _) (a : _),
|
||||
R _ _ _ _ H a
|
||||
|
||||
print Environment 1
|
||||
print environment 1
|
||||
|
|
|
@ -4,9 +4,9 @@
|
|||
Assumed: D
|
||||
Assumed: R
|
||||
Proved: R2
|
||||
Theorem R2 (A A' B B' : Type) (H : (A → B) = (A' → B')) (a : A) : B = B' := R A A' (λ x : A, B) (λ x : A', B') H a
|
||||
theorem R2 (A A' B B' : Type) (H : (A → B) = (A' → B')) (a : A) : B = B' := R A A' (λ x : A, B) (λ x : A', B') H a
|
||||
Proved: R3
|
||||
Proved: R4
|
||||
Proved: R5
|
||||
Theorem R5 (A1 A2 B1 B2 : Type) (H : (A1 → B1) = (A2 → B2)) (a : A1) : B1 = B2 :=
|
||||
theorem R5 (A1 A2 B1 B2 : Type) (H : (A1 → B1) = (A2 → B2)) (a : A1) : B1 = B2 :=
|
||||
R A1 A2 (λ x : A1, B1) (λ x : A2, B2) H a
|
||||
|
|
|
@ -1,12 +1,12 @@
|
|||
Variable C : Pi (A B : Type) (H : A = B) (a : A), B
|
||||
variable C : Pi (A B : Type) (H : A = B) (a : A), B
|
||||
|
||||
Variable D : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)), A = A'
|
||||
variable D : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)), A = A'
|
||||
|
||||
Variable R : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A),
|
||||
variable R : Pi (A A' : Type) (B : A -> Type) (B' : A' -> Type) (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A),
|
||||
(B a) = (B' (C A A' (D A A' B B' H) a))
|
||||
|
||||
Theorem R2 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
theorem R2 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
fun A1 A2 B1 B2 H a,
|
||||
R _ _ _ _ H a
|
||||
|
||||
print Environment 1.
|
||||
print environment 1.
|
|
@ -4,5 +4,5 @@
|
|||
Assumed: D
|
||||
Assumed: R
|
||||
Proved: R2
|
||||
Theorem R2 (A1 A2 B1 B2 : Type) (H : (A1 → B1) = (A2 → B2)) (a : A1) : B1 = B2 :=
|
||||
theorem R2 (A1 A2 B1 B2 : Type) (H : (A1 → B1) = (A2 → B2)) (a : A1) : B1 = B2 :=
|
||||
R A1 A2 (λ x : A1, B1) (λ x : A2, B2) H a
|
||||
|
|
|
@ -1,12 +1,12 @@
|
|||
Variable C {A B : Type} (H : A = B) (a : A) : B
|
||||
variable C {A B : Type} (H : A = B) (a : A) : B
|
||||
|
||||
Variable D {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) : A = A'
|
||||
variable D {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) : A = A'
|
||||
|
||||
Variable R {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A) :
|
||||
variable R {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A) :
|
||||
(B a) = (B' (C (D H) a))
|
||||
|
||||
Theorem R2 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
theorem R2 : Pi (A1 A2 B1 B2 : Type) (H : (A1 -> B1) = (A2 -> B2)) (a : A1), B1 = B2 :=
|
||||
fun A1 A2 B1 B2 H a, R H a
|
||||
|
||||
SetOption pp::implicit true
|
||||
print Environment 7.
|
||||
setoption pp::implicit true
|
||||
print environment 7.
|
||||
|
|
|
@ -5,12 +5,12 @@
|
|||
Assumed: R
|
||||
Proved: R2
|
||||
Set: lean::pp::implicit
|
||||
Import "kernel"
|
||||
Import "Nat"
|
||||
Variable C {A B : Type} (H : @eq Type A B) (a : A) : B
|
||||
Variable D {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) :
|
||||
import "kernel"
|
||||
import "Nat"
|
||||
variable C {A B : Type} (H : @eq Type A B) (a : A) : B
|
||||
variable D {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) :
|
||||
@eq Type A A'
|
||||
Variable R {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) (a : A) :
|
||||
variable R {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) (a : A) :
|
||||
@eq Type (B a) (B' (@C A A' (@D A A' (λ x : A, B x) (λ x : A', B' x) H) a))
|
||||
Theorem R2 (A1 A2 B1 B2 : Type) (H : @eq Type (A1 → B1) (A2 → B2)) (a : A1) : @eq Type B1 B2 :=
|
||||
theorem R2 (A1 A2 B1 B2 : Type) (H : @eq Type (A1 → B1) (A2 → B2)) (a : A1) : @eq Type B1 B2 :=
|
||||
@R A1 A2 (λ x : A1, B1) (λ x : A2, B2) H a
|
||||
|
|
|
@ -1,12 +1,12 @@
|
|||
Variable C {A B : Type} (H : A = B) (a : A) : B
|
||||
variable C {A B : Type} (H : A = B) (a : A) : B
|
||||
|
||||
Variable D {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) : A = A'
|
||||
variable D {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) : A = A'
|
||||
|
||||
Variable R {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A) :
|
||||
variable R {A A' : Type} {B : A -> Type} {B' : A' -> Type} (H : (Pi x : A, B x) = (Pi x : A', B' x)) (a : A) :
|
||||
(B a) = (B' (C (D H) a))
|
||||
|
||||
Theorem R2 : Pi (A1 A2 B1 B2 : Type), ((A1 -> B1) = (A2 -> B2)) -> A1 -> (B1 = B2) :=
|
||||
theorem R2 : Pi (A1 A2 B1 B2 : Type), ((A1 -> B1) = (A2 -> B2)) -> A1 -> (B1 = B2) :=
|
||||
fun A1 A2 B1 B2 H a, R H a
|
||||
|
||||
SetOption pp::implicit true
|
||||
print Environment 7.
|
||||
setoption pp::implicit true
|
||||
print environment 7.
|
||||
|
|
|
@ -5,12 +5,12 @@
|
|||
Assumed: R
|
||||
Proved: R2
|
||||
Set: lean::pp::implicit
|
||||
Import "kernel"
|
||||
Import "Nat"
|
||||
Variable C {A B : Type} (H : @eq Type A B) (a : A) : B
|
||||
Variable D {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) :
|
||||
import "kernel"
|
||||
import "Nat"
|
||||
variable C {A B : Type} (H : @eq Type A B) (a : A) : B
|
||||
variable D {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) :
|
||||
@eq Type A A'
|
||||
Variable R {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) (a : A) :
|
||||
variable R {A A' : Type} {B : A → Type} {B' : A' → Type} (H : @eq Type (Π x : A, B x) (Π x : A', B' x)) (a : A) :
|
||||
@eq Type (B a) (B' (@C A A' (@D A A' (λ x : A, B x) (λ x : A', B' x) H) a))
|
||||
Theorem R2 (A1 A2 B1 B2 : Type) (H : @eq Type (A1 → B1) (A2 → B2)) (a : A1) : @eq Type B1 B2 :=
|
||||
theorem R2 (A1 A2 B1 B2 : Type) (H : @eq Type (A1 → B1) (A2 → B2)) (a : A1) : @eq Type B1 B2 :=
|
||||
@R A1 A2 (λ x : A1, B1) (λ x : A2, B2) H a
|
||||
|
|
|
@ -1,21 +1,21 @@
|
|||
Variables A B C : (Type U)
|
||||
Variable P : A -> Bool
|
||||
Variable F1 : A -> B -> C
|
||||
Variable F2 : A -> B -> C
|
||||
Variable H : Pi (a : A) (b : B), (F1 a b) == (F2 a b)
|
||||
Variable a : A
|
||||
Check Eta (F2 a)
|
||||
Check Abst (fun a : A,
|
||||
variables A B C : (Type U)
|
||||
variable P : A -> Bool
|
||||
variable F1 : A -> B -> C
|
||||
variable F2 : A -> B -> C
|
||||
variable H : Pi (a : A) (b : B), (F1 a b) == (F2 a b)
|
||||
variable a : A
|
||||
check Eta (F2 a)
|
||||
check Abst (fun a : A,
|
||||
(Trans (Symm (Eta (F1 a)))
|
||||
(Trans (Abst (fun (b : B), H a b))
|
||||
(Eta (F2 a)))))
|
||||
|
||||
Check Abst (fun a, (Abst (fun b, H a b)))
|
||||
check Abst (fun a, (Abst (fun b, H a b)))
|
||||
|
||||
Theorem T1 : F1 = F2 := Abst (fun a, (Abst (fun b, H a b)))
|
||||
Theorem T2 : (fun (x1 : A) (x2 : B), F1 x1 x2) = F2 := Abst (fun a, (Abst (fun b, H a b)))
|
||||
Theorem T3 : F1 = (fun (x1 : A) (x2 : B), F2 x1 x2) := Abst (fun a, (Abst (fun b, H a b)))
|
||||
Theorem T4 : (fun (x1 : A) (x2 : B), F1 x1 x2) = (fun (x1 : A) (x2 : B), F2 x1 x2) := Abst (fun a, (Abst (fun b, H a b)))
|
||||
print Environment 4
|
||||
SetOption pp::implicit true
|
||||
print Environment 4
|
||||
theorem T1 : F1 = F2 := Abst (fun a, (Abst (fun b, H a b)))
|
||||
theorem T2 : (fun (x1 : A) (x2 : B), F1 x1 x2) = F2 := Abst (fun a, (Abst (fun b, H a b)))
|
||||
theorem T3 : F1 = (fun (x1 : A) (x2 : B), F2 x1 x2) := Abst (fun a, (Abst (fun b, H a b)))
|
||||
theorem T4 : (fun (x1 : A) (x2 : B), F1 x1 x2) = (fun (x1 : A) (x2 : B), F2 x1 x2) := Abst (fun a, (Abst (fun b, H a b)))
|
||||
print environment 4
|
||||
setoption pp::implicit true
|
||||
print environment 4
|
|
@ -16,27 +16,27 @@ Abst (λ a : A, Abst (λ b : B, H a b)) : (λ (x : A) (x::1 : B), F1 x x::1) ==
|
|||
Proved: T2
|
||||
Proved: T3
|
||||
Proved: T4
|
||||
Theorem T1 : F1 = F2 := Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
Theorem T2 : (λ (x1 : A) (x2 : B), F1 x1 x2) = F2 := Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
Theorem T3 : F1 = (λ (x1 : A) (x2 : B), F2 x1 x2) := Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
Theorem T4 : (λ (x1 : A) (x2 : B), F1 x1 x2) = (λ (x1 : A) (x2 : B), F2 x1 x2) :=
|
||||
theorem T1 : F1 = F2 := Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
theorem T2 : (λ (x1 : A) (x2 : B), F1 x1 x2) = F2 := Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
theorem T3 : F1 = (λ (x1 : A) (x2 : B), F2 x1 x2) := Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
theorem T4 : (λ (x1 : A) (x2 : B), F1 x1 x2) = (λ (x1 : A) (x2 : B), F2 x1 x2) :=
|
||||
Abst (λ a : A, Abst (λ b : B, H a b))
|
||||
Set: lean::pp::implicit
|
||||
Theorem T1 : @eq (A → B → C) F1 F2 :=
|
||||
theorem T1 : @eq (A → B → C) F1 F2 :=
|
||||
@Abst A (λ x : A, B → C) F1 F2 (λ a : A, @Abst B (λ x : B, C) (F1 a) (F2 a) (λ b : B, H a b))
|
||||
Theorem T2 : @eq (A → B → C) (λ (x1 : A) (x2 : B), F1 x1 x2) F2 :=
|
||||
theorem T2 : @eq (A → B → C) (λ (x1 : A) (x2 : B), F1 x1 x2) F2 :=
|
||||
@Abst A
|
||||
(λ x : A, B → C)
|
||||
(λ (x1 : A) (x2 : B), F1 x1 x2)
|
||||
F2
|
||||
(λ a : A, @Abst B (λ x : B, C) (λ x2 : B, F1 a x2) (F2 a) (λ b : B, H a b))
|
||||
Theorem T3 : @eq (A → B → C) F1 (λ (x1 : A) (x2 : B), F2 x1 x2) :=
|
||||
theorem T3 : @eq (A → B → C) F1 (λ (x1 : A) (x2 : B), F2 x1 x2) :=
|
||||
@Abst A
|
||||
(λ x : A, B → C)
|
||||
F1
|
||||
(λ (x1 : A) (x2 : B), F2 x1 x2)
|
||||
(λ a : A, @Abst B (λ x : B, C) (F1 a) (λ x2 : B, F2 a x2) (λ b : B, H a b))
|
||||
Theorem T4 : @eq (A → B → C) (λ (x1 : A) (x2 : B), F1 x1 x2) (λ (x1 : A) (x2 : B), F2 x1 x2) :=
|
||||
theorem T4 : @eq (A → B → C) (λ (x1 : A) (x2 : B), F1 x1 x2) (λ (x1 : A) (x2 : B), F2 x1 x2) :=
|
||||
@Abst A
|
||||
(λ x : A, B → C)
|
||||
(λ (x1 : A) (x2 : B), F1 x1 x2)
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
Import Int.
|
||||
Variable i : Int
|
||||
Check i = 0
|
||||
SetOption pp::coercion true
|
||||
Check i = 0
|
||||
import Int.
|
||||
variable i : Int
|
||||
check i = 0
|
||||
setoption pp::coercion true
|
||||
check i = 0
|
||||
|
|
|
@ -1,8 +1,8 @@
|
|||
Import Int.
|
||||
Variable List : Type -> Type
|
||||
Variable nil {A : Type} : List A
|
||||
Variable cons {A : Type} (head : A) (tail : List A) : List A
|
||||
Variable l : List Int.
|
||||
Check l = nil.
|
||||
SetOption pp::implicit true
|
||||
Check l = nil.
|
||||
import Int.
|
||||
variable List : Type -> Type
|
||||
variable nil {A : Type} : List A
|
||||
variable cons {A : Type} (head : A) (tail : List A) : List A
|
||||
variable l : List Int.
|
||||
check l = nil.
|
||||
setoption pp::implicit true
|
||||
check l = nil.
|
||||
|
|
|
@ -1,10 +1,10 @@
|
|||
Variable Vector : Nat -> Type
|
||||
Variable n : Nat
|
||||
Variable v1 : Vector n
|
||||
Variable v2 : Vector (n + 0)
|
||||
Variable v3 : Vector (0 + n)
|
||||
Axiom H1 : v1 == v2
|
||||
Axiom H2 : v2 == v3
|
||||
Check HTrans H1 H2
|
||||
SetOption pp::implicit true
|
||||
Check HTrans H1 H2
|
||||
variable Vector : Nat -> Type
|
||||
variable n : Nat
|
||||
variable v1 : Vector n
|
||||
variable v2 : Vector (n + 0)
|
||||
variable v3 : Vector (0 + n)
|
||||
axiom H1 : v1 == v2
|
||||
axiom H2 : v2 == v3
|
||||
check HTrans H1 H2
|
||||
setoption pp::implicit true
|
||||
check HTrans H1 H2
|
||||
|
|
|
@ -1,8 +1,8 @@
|
|||
Eval fun x, x
|
||||
eval fun x, x
|
||||
print fun x, x
|
||||
|
||||
Check fun x, x
|
||||
Theorem T (A : Type) (x : A) : Pi (y : A), A
|
||||
check fun x, x
|
||||
theorem T (A : Type) (x : A) : Pi (y : A), A
|
||||
:= _.
|
||||
|
||||
Theorem T (x : _) : x = x := Refl x.
|
||||
theorem T (x : _) : x = x := Refl x.
|
||||
|
|
|
@ -1,18 +1,18 @@
|
|||
-- comment
|
||||
print true
|
||||
SetOption lean::pp::notation false
|
||||
setoption lean::pp::notation false
|
||||
print true && false
|
||||
SetOption pp::unicode false
|
||||
setoption pp::unicode false
|
||||
print true && false
|
||||
Variable a : Bool
|
||||
Variable a : Bool
|
||||
Variable b : Bool
|
||||
variable a : Bool
|
||||
variable a : Bool
|
||||
variable b : Bool
|
||||
print a && b
|
||||
Variable A : Type
|
||||
Check a && A
|
||||
print Environment 1
|
||||
print Options
|
||||
SetOption lean::p::notation true
|
||||
SetOption lean::pp::notation 10
|
||||
SetOption lean::pp::notation true
|
||||
variable A : Type
|
||||
check a && A
|
||||
print environment 1
|
||||
print options
|
||||
setoption lean::p::notation true
|
||||
setoption lean::pp::notation 10
|
||||
setoption lean::pp::notation true
|
||||
print a && b
|
||||
|
|
|
@ -17,7 +17,7 @@ Function type:
|
|||
Arguments types:
|
||||
a : Bool
|
||||
A : Type
|
||||
Variable A : Type
|
||||
variable A : Type
|
||||
(lean::pp::notation := false, pp::unicode := false, pp::colors := false)
|
||||
Error (line: 15, pos: 10) unknown option 'lean::p::notation', type 'Help Options.' for list of available options
|
||||
Error (line: 16, pos: 29) invalid option value, given option is not an integer
|
||||
|
|
|
@ -1,9 +1,9 @@
|
|||
Variable myeq : Pi (A : Type), A -> A -> Bool
|
||||
variable myeq : Pi (A : Type), A -> A -> Bool
|
||||
print myeq _ true false
|
||||
Variable T : Type
|
||||
Variable a : T
|
||||
Check myeq _ true a
|
||||
Variable myeq2 {A:Type} (a b : A) : Bool
|
||||
Infix 50 === : myeq2
|
||||
SetOption lean::pp::implicit true
|
||||
Check true === a
|
||||
variable T : Type
|
||||
variable a : T
|
||||
check myeq _ true a
|
||||
variable myeq2 {A:Type} (a b : A) : Bool
|
||||
infix 50 === : myeq2
|
||||
setoption lean::pp::implicit true
|
||||
check true === a
|
|
@ -1,6 +1,6 @@
|
|||
Import Int.
|
||||
Variable a : Int
|
||||
Variable P : Int -> Int -> Bool
|
||||
Axiom H : P a a
|
||||
Theorem T : exists x : Int, P a a := ExistsIntro a H.
|
||||
print Environment 1.
|
||||
import Int.
|
||||
variable a : Int
|
||||
variable P : Int -> Int -> Bool
|
||||
axiom H : P a a
|
||||
theorem T : exists x : Int, P a a := ExistsIntro a H.
|
||||
print environment 1.
|
|
@ -5,4 +5,4 @@
|
|||
Assumed: P
|
||||
Assumed: H
|
||||
Proved: T
|
||||
Theorem T : ∃ x : ℤ, P a a := ExistsIntro a H
|
||||
theorem T : ∃ x : ℤ, P a a := ExistsIntro a H
|
||||
|
|
|
@ -1,18 +1,18 @@
|
|||
Import Int.
|
||||
Variable a : Int
|
||||
Variable P : Int -> Int -> Bool
|
||||
Variable f : Int -> Int -> Int
|
||||
Variable g : Int -> Int
|
||||
Axiom H1 : P (f a (g a)) (f a (g a))
|
||||
Axiom H2 : P (f (g a) (g a)) (f (g a) (g a))
|
||||
Axiom H3 : P (f (g a) (g a)) (g a)
|
||||
Theorem T1 : exists x y : Int, P (f y x) (f y x) := ExistsIntro _ (ExistsIntro _ H1)
|
||||
Theorem T2 : exists x : Int, P (f x (g x)) (f x (g x)) := ExistsIntro _ H1
|
||||
Theorem T3 : exists x : Int, P (f x x) (f x x) := ExistsIntro _ H2
|
||||
Theorem T4 : exists x : Int, P (f (g a) x) (f x x) := ExistsIntro _ H2
|
||||
Theorem T5 : exists x : Int, P x x := ExistsIntro _ H2
|
||||
Theorem T6 : exists x y : Int, P x y := ExistsIntro _ (ExistsIntro _ H3)
|
||||
Theorem T7 : exists x : Int, P (f x x) x := ExistsIntro _ H3
|
||||
Theorem T8 : exists x y : Int, P (f x x) y := ExistsIntro _ (ExistsIntro _ H3)
|
||||
import Int.
|
||||
variable a : Int
|
||||
variable P : Int -> Int -> Bool
|
||||
variable f : Int -> Int -> Int
|
||||
variable g : Int -> Int
|
||||
axiom H1 : P (f a (g a)) (f a (g a))
|
||||
axiom H2 : P (f (g a) (g a)) (f (g a) (g a))
|
||||
axiom H3 : P (f (g a) (g a)) (g a)
|
||||
theorem T1 : exists x y : Int, P (f y x) (f y x) := ExistsIntro _ (ExistsIntro _ H1)
|
||||
theorem T2 : exists x : Int, P (f x (g x)) (f x (g x)) := ExistsIntro _ H1
|
||||
theorem T3 : exists x : Int, P (f x x) (f x x) := ExistsIntro _ H2
|
||||
theorem T4 : exists x : Int, P (f (g a) x) (f x x) := ExistsIntro _ H2
|
||||
theorem T5 : exists x : Int, P x x := ExistsIntro _ H2
|
||||
theorem T6 : exists x y : Int, P x y := ExistsIntro _ (ExistsIntro _ H3)
|
||||
theorem T7 : exists x : Int, P (f x x) x := ExistsIntro _ H3
|
||||
theorem T8 : exists x y : Int, P (f x x) y := ExistsIntro _ (ExistsIntro _ H3)
|
||||
|
||||
print Environment 8.
|
||||
print environment 8.
|
Some files were not shown because too many files have changed in this diff Show more
Loading…
Reference in a new issue