diff --git a/examples/lean/even.lean b/examples/lean/even.lean index b4a4ba733..f76c20bec 100644 --- a/examples/lean/even.lean +++ b/examples/lean/even.lean @@ -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. diff --git a/examples/lean/ex1.lean b/examples/lean/ex1.lean index dff6e1ad5..8e14b78cb 100644 --- a/examples/lean/ex1.lean +++ b/examples/lean/ex1.lean @@ -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 diff --git a/examples/lean/ex2.lean b/examples/lean/ex2.lean index 06a1ea9b2..5a0e3a0c1 100644 --- a/examples/lean/ex2.lean +++ b/examples/lean/ex2.lean @@ -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. diff --git a/examples/lean/ex3.lean b/examples/lean/ex3.lean index 44f8ef87e..d5d8aec11 100644 --- a/examples/lean/ex3.lean +++ b/examples/lean/ex3.lean @@ -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. \ No newline at end of file +print environment 3. \ No newline at end of file diff --git a/examples/lean/ex4.lean b/examples/lean/ex4.lean index 48521901d..23bcde8d6 100644 --- a/examples/lean/ex4.lean +++ b/examples/lean/ex4.lean @@ -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 diff --git a/examples/lean/ex5.lean b/examples/lean/ex5.lean index 20ebb895d..b2866682c 100644 --- a/examples/lean/ex5.lean +++ b/examples/lean/ex5.lean @@ -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 \ No newline at end of file +print environment 2 \ No newline at end of file diff --git a/examples/lean/set.lean b/examples/lean/set.lean index 70d518cc7..1ebbc059c 100644 --- a/examples/lean/set.lean +++ b/examples/lean/set.lean @@ -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)) diff --git a/examples/lean/tactic1.lean b/examples/lean/tactic1.lean index fe07510ca..4de6d4bbb 100644 --- a/examples/lean/tactic1.lean +++ b/examples/lean/tactic1.lean @@ -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 diff --git a/examples/lean/tactic_in_lua.lean b/examples/lean/tactic_in_lua.lean index b6c963ba0..c3d10e86d 100644 --- a/examples/lean/tactic_in_lua.lean +++ b/examples/lean/tactic_in_lua.lean @@ -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. diff --git a/src/builtin/Int.lean b/src/builtin/Int.lean index 87b3c6687..fe7bcc9c6 100644 --- a/src/builtin/Int.lean +++ b/src/builtin/Int.lean @@ -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. \ No newline at end of file +setopaque sub true +setopaque neg true +end \ No newline at end of file diff --git a/src/builtin/Nat.lean b/src/builtin/Nat.lean index db39a8548..6e2e3feaa 100644 --- a/src/builtin/Nat.lean +++ b/src/builtin/Nat.lean @@ -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. \ No newline at end of file +setopaque ge true +setopaque lt true +setopaque gt true +setopaque id true +end \ No newline at end of file diff --git a/src/builtin/Real.lean b/src/builtin/Real.lean index b6d4b8d1a..daf8585f6 100644 --- a/src/builtin/Real.lean +++ b/src/builtin/Real.lean @@ -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 diff --git a/src/builtin/cast.lean b/src/builtin/cast.lean index c440b311f..6d9813024 100644 --- a/src/builtin/cast.lean +++ b/src/builtin/cast.lean @@ -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). diff --git a/src/builtin/find.lua b/src/builtin/find.lua index bf4739e09..8814ec37e 100644 --- a/src/builtin/find.lua +++ b/src/builtin/find.lua @@ -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() diff --git a/src/builtin/kernel.lean b/src/builtin/kernel.lean index cf6917186..92cfbbe21 100644 --- a/src/builtin/kernel.lean +++ b/src/builtin/kernel.lean @@ -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 diff --git a/src/builtin/obj/Int.olean b/src/builtin/obj/Int.olean index 0f9ddad3e..0f6d1e7a7 100644 Binary files a/src/builtin/obj/Int.olean and b/src/builtin/obj/Int.olean differ diff --git a/src/builtin/obj/Nat.olean b/src/builtin/obj/Nat.olean index da8b9bc90..5da36280d 100644 Binary files a/src/builtin/obj/Nat.olean and b/src/builtin/obj/Nat.olean differ diff --git a/src/builtin/obj/Real.olean b/src/builtin/obj/Real.olean index c095bf86a..6035111cf 100644 Binary files a/src/builtin/obj/Real.olean and b/src/builtin/obj/Real.olean differ diff --git a/src/builtin/obj/cast.olean b/src/builtin/obj/cast.olean index 9c513c929..52eac2499 100644 Binary files a/src/builtin/obj/cast.olean and b/src/builtin/obj/cast.olean differ diff --git a/src/builtin/obj/kernel.olean b/src/builtin/obj/kernel.olean index 0f2b18280..fb4339ac3 100644 Binary files a/src/builtin/obj/kernel.olean and b/src/builtin/obj/kernel.olean differ diff --git a/src/builtin/obj/specialfn.olean b/src/builtin/obj/specialfn.olean index 4862612d7..46af94c19 100644 Binary files a/src/builtin/obj/specialfn.olean and b/src/builtin/obj/specialfn.olean differ diff --git a/src/builtin/specialfn.lean b/src/builtin/specialfn.lean index 98b72613e..95f4e4e71 100644 --- a/src/builtin/specialfn.lean +++ b/src/builtin/specialfn.lean @@ -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. diff --git a/src/frontends/lean/coercion.h b/src/frontends/lean/coercion.h index addd096c9..666349d2f 100644 --- a/src/frontends/lean/coercion.h +++ b/src/frontends/lean/coercion.h @@ -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; }; diff --git a/src/frontends/lean/operator_info.cpp b/src/frontends/lean/operator_info.cpp index 9d8dba01c..ee2c4a08c 100644 --- a/src/frontends/lean/operator_info.cpp +++ b/src/frontends/lean/operator_info.cpp @@ -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); } diff --git a/src/frontends/lean/parser_cmds.cpp b/src/frontends/lean/parser_cmds.cpp index 750dcb482..28143a74c 100644 --- a/src/frontends/lean/parser_cmds.cpp +++ b/src/frontends/lean/parser_cmds.cpp @@ -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 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 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 { diff --git a/src/frontends/lean/parser_imp.h b/src/frontends/lean/parser_imp.h index 589dfa4d8..2f22729fb 100644 --- a/src/frontends/lean/parser_imp.h +++ b/src/frontends/lean/parser_imp.h @@ -73,6 +73,8 @@ class parser_imp { pos_info m_last_script_pos; tactic_hints m_tactic_hints; std::vector m_namespace_prefixes; + enum class scope_kind { Scope, Namespace }; + std::vector m_scope_kinds; std::unique_ptr 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(); /*@}*/ diff --git a/src/kernel/environment.cpp b/src/kernel/environment.cpp index e96e3e21d..ae6e26694 100644 --- a/src/kernel/environment.cpp +++ b/src/kernel/environment.cpp @@ -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: diff --git a/src/kernel/object.cpp b/src/kernel/object.cpp index 816517439..e701b79ab 100644 --- a/src/kernel/object.cpp +++ b/src/kernel/object.cpp @@ -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)); } diff --git a/src/tests/frontends/lean/parser.cpp b/src/tests/frontends/lean/parser.cpp index df7c9c3cf..b628fe7ee 100644 --- a/src/tests/frontends/lean/parser.cpp +++ b/src/tests/frontends/lean/parser.cpp @@ -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"); } diff --git a/src/tests/library/formatter.cpp b/src/tests/library/formatter.cpp index 83447efd2..7e55baa78 100644 --- a/src/tests/library/formatter.cpp +++ b/src/tests/library/formatter.cpp @@ -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)); diff --git a/tests/lean/abst.lean b/tests/lean/abst.lean index fb5c960e2..1c0fcd506 100644 --- a/tests/lean/abst.lean +++ b/tests/lean/abst.lean @@ -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))). diff --git a/tests/lean/alias1.lean b/tests/lean/alias1.lean index ee5094e78..8b07b25a1 100644 --- a/tests/lean/alias1.lean +++ b/tests/lean/alias1.lean @@ -1,3 +1,3 @@ -Alias BB : Bool. -Variable x : BB. -print Environment 1. \ No newline at end of file +alias BB : Bool. +variable x : BB. +print environment 1. \ No newline at end of file diff --git a/tests/lean/alias1.lean.expected.out b/tests/lean/alias1.lean.expected.out index 89366d1bc..929fb1017 100644 --- a/tests/lean/alias1.lean.expected.out +++ b/tests/lean/alias1.lean.expected.out @@ -1,4 +1,4 @@ Set: pp::colors Set: pp::unicode Assumed: x -Variable x : BB +variable x : BB diff --git a/tests/lean/alias2.lean b/tests/lean/alias2.lean index bfbc23453..008602499 100644 --- a/tests/lean/alias2.lean +++ b/tests/lean/alias2.lean @@ -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 \ No newline at end of file +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 \ No newline at end of file diff --git a/tests/lean/alias2.lean.expected.out b/tests/lean/alias2.lean.expected.out index c49d6dbe4..1b4b53b26 100644 --- a/tests/lean/alias2.lean.expected.out +++ b/tests/lean/alias2.lean.expected.out @@ -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 diff --git a/tests/lean/apply_tac1.lean b/tests/lean/apply_tac1.lean index 950b0e68d..e7022ea00 100644 --- a/tests/lean/apply_tac1.lean +++ b/tests/lean/apply_tac1.lean @@ -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. \ No newline at end of file +print environment 2. \ No newline at end of file diff --git a/tests/lean/apply_tac1.lean.expected.out b/tests/lean/apply_tac1.lean.expected.out index bd09665a2..d184b1342 100644 --- a/tests/lean/apply_tac1.lean.expected.out +++ b/tests/lean/apply_tac1.lean.expected.out @@ -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) diff --git a/tests/lean/apply_tac2.lean b/tests/lean/apply_tac2.lean index 7e3c36321..f77fc8d5b 100644 --- a/tests/lean/apply_tac2.lean +++ b/tests/lean/apply_tac2.lean @@ -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. diff --git a/tests/lean/arith1.lean b/tests/lean/arith1.lean index aed715373..c9b844932 100644 --- a/tests/lean/arith1.lean +++ b/tests/lean/arith1.lean @@ -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 diff --git a/tests/lean/arith2.lean b/tests/lean/arith2.lean index 656d50e27..d2adaa403 100644 --- a/tests/lean/arith2.lean +++ b/tests/lean/arith2.lean @@ -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 diff --git a/tests/lean/arith3.lean b/tests/lean/arith3.lean index 58f055fff..898bc60f6 100644 --- a/tests/lean/arith3.lean +++ b/tests/lean/arith3.lean @@ -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) \ No newline at end of file +eval -8 mod 3 +eval (-8 div 3)*3 + (-8 mod 3) \ No newline at end of file diff --git a/tests/lean/arith4.lean b/tests/lean/arith4.lean index e20c7bc3e..93d3c315a 100644 --- a/tests/lean/arith4.lean +++ b/tests/lean/arith4.lean @@ -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) \ No newline at end of file +import specialfn. +variable x : Real +eval sin(x) +eval cos(x) +eval tan(x) +eval cot(x) +eval sec(x) +eval csc(x) \ No newline at end of file diff --git a/tests/lean/arith5.lean b/tests/lean/arith5.lean index 77ccb0495..3fab16b9b 100644 --- a/tests/lean/arith5.lean +++ b/tests/lean/arith5.lean @@ -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) \ No newline at end of file +import specialfn. +variable x : Real +eval sinh(x) +eval cosh(x) +eval tanh(x) +eval coth(x) +eval sech(x) +eval csch(x) \ No newline at end of file diff --git a/tests/lean/arith6.lean b/tests/lean/arith6.lean index 7ab9461ce..5f5ced6e8 100644 --- a/tests/lean/arith6.lean +++ b/tests/lean/arith6.lean @@ -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 \ No newline at end of file diff --git a/tests/lean/arith7.lean b/tests/lean/arith7.lean index aa9c1cf17..e1f749316 100644 --- a/tests/lean/arith7.lean +++ b/tests/lean/arith7.lean @@ -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 \ No newline at end of file diff --git a/tests/lean/arith8.lean b/tests/lean/arith8.lean index f1ed8452b..20f92aa96 100644 --- a/tests/lean/arith8.lean +++ b/tests/lean/arith8.lean @@ -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 \ No newline at end of file +import Real. +eval 10.3 +eval 0.3 +eval 0.3 + 0.1 +eval 0.2 + 0.7 +eval 1/3 + 0.1 \ No newline at end of file diff --git a/tests/lean/arrow.lean b/tests/lean/arrow.lean index d575815bc..5b7d24671 100644 --- a/tests/lean/arrow.lean +++ b/tests/lean/arrow.lean @@ -1,4 +1,4 @@ -Import Int. +import Int. print (Int -> Int) -> Int print Int -> Int -> Int print Int -> (Int -> Int) diff --git a/tests/lean/bad1.lean b/tests/lean/bad1.lean index 4905f5e64..56070772a 100644 --- a/tests/lean/bad1.lean +++ b/tests/lean/bad1.lean @@ -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 diff --git a/tests/lean/bad10.lean b/tests/lean/bad10.lean index 3a5d58bf1..675736312 100644 --- a/tests/lean/bad10.lean +++ b/tests/lean/bad10.lean @@ -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. diff --git a/tests/lean/bad10.lean.expected.out b/tests/lean/bad10.lean.expected.out index 0a75399dd..36e79f708 100644 --- a/tests/lean/bad10.lean.expected.out +++ b/tests/lean/bad10.lean.expected.out @@ -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 diff --git a/tests/lean/bad2.lean b/tests/lean/bad2.lean index d93955ee9..754f798f4 100644 --- a/tests/lean/bad2.lean +++ b/tests/lean/bad2.lean @@ -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. \ No newline at end of file +setoption pp::coercion true +setoption pp::implicit true +print environment 1. \ No newline at end of file diff --git a/tests/lean/bad2.lean.expected.out b/tests/lean/bad2.lean.expected.out index 45e2c0671..8186a7422 100644 --- a/tests/lean/bad2.lean.expected.out +++ b/tests/lean/bad2.lean.expected.out @@ -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 ℕ) diff --git a/tests/lean/bad3.lean b/tests/lean/bad3.lean index 6d3d21afd..f47cdcdfc 100644 --- a/tests/lean/bad3.lean +++ b/tests/lean/bad3.lean @@ -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)) diff --git a/tests/lean/bad4.lean b/tests/lean/bad4.lean index f1fe106e2..9a26be3ed 100644 --- a/tests/lean/bad4.lean +++ b/tests/lean/bad4.lean @@ -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 diff --git a/tests/lean/bad5.lean b/tests/lean/bad5.lean index 56ea5fcab..017e71f63 100644 --- a/tests/lean/bad5.lean +++ b/tests/lean/bad5.lean @@ -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 diff --git a/tests/lean/bad5.lean.expected.out b/tests/lean/bad5.lean.expected.out index b87f36834..deea5aa7d 100644 --- a/tests/lean/bad5.lean.expected.out +++ b/tests/lean/bad5.lean.expected.out @@ -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 diff --git a/tests/lean/bad6.lean b/tests/lean/bad6.lean index 72a85ac44..1a026d510 100644 --- a/tests/lean/bad6.lean +++ b/tests/lean/bad6.lean @@ -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 diff --git a/tests/lean/bad7.lean b/tests/lean/bad7.lean index 472edbf37..8b015e25d 100644 --- a/tests/lean/bad7.lean +++ b/tests/lean/bad7.lean @@ -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 diff --git a/tests/lean/bad7.lean.expected.out b/tests/lean/bad7.lean.expected.out index a72050168..2ff82c14c 100644 --- a/tests/lean/bad7.lean.expected.out +++ b/tests/lean/bad7.lean.expected.out @@ -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 diff --git a/tests/lean/bad8.lean b/tests/lean/bad8.lean index 91fc0671c..b85afb4d7 100644 --- a/tests/lean/bad8.lean +++ b/tests/lean/bad8.lean @@ -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)) diff --git a/tests/lean/bad9.lean b/tests/lean/bad9.lean index 2b8eab1c8..6d72f6d12 100644 --- a/tests/lean/bad9.lean +++ b/tests/lean/bad9.lean @@ -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. \ No newline at end of file diff --git a/tests/lean/bug.lean b/tests/lean/bug.lean index b076df7cc..6e093e983 100644 --- a/tests/lean/bug.lean +++ b/tests/lean/bug.lean @@ -1 +1 @@ -Check fun (A A' : (Type U)) (H : A == A'), Symm H +check fun (A A' : (Type U)) (H : A == A'), Symm H diff --git a/tests/lean/calc1.lean b/tests/lean/calc1.lean index 8e579e8ba..2f9412341 100644 --- a/tests/lean/calc1.lean +++ b/tests/lean/calc1.lean @@ -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) diff --git a/tests/lean/calc2.lean b/tests/lean/calc2.lean index 178fb6539..b759371c0 100644 --- a/tests/lean/calc2.lean +++ b/tests/lean/calc2.lean @@ -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 }. diff --git a/tests/lean/cast1.lean b/tests/lean/cast1.lean index 0cc99532d..33a9d76e8 100644 --- a/tests/lean/cast1.lean +++ b/tests/lean/cast1.lean @@ -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) diff --git a/tests/lean/cast2.lean b/tests/lean/cast2.lean index 91c98d7e2..4bedf9335 100644 --- a/tests/lean/cast2.lean +++ b/tests/lean/cast2.lean @@ -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 diff --git a/tests/lean/cast3.lean b/tests/lean/cast3.lean index 862633897..4d29fd438 100644 --- a/tests/lean/cast3.lean +++ b/tests/lean/cast3.lean @@ -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)) diff --git a/tests/lean/cast4.lean b/tests/lean/cast4.lean index 4b3332758..f26a279d3 100644 --- a/tests/lean/cast4.lean +++ b/tests/lean/cast4.lean @@ -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) diff --git a/tests/lean/coercion1.lean b/tests/lean/coercion1.lean index 778860103..6814fbeb5 100644 --- a/tests/lean/coercion1.lean +++ b/tests/lean/coercion1.lean @@ -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 diff --git a/tests/lean/coercion1.lean.expected.out b/tests/lean/coercion1.lean.expected.out index 1a686fa4d..27689cbb1 100644 --- a/tests/lean/coercion1.lean.expected.out +++ b/tests/lean/coercion1.lean.expected.out @@ -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 diff --git a/tests/lean/coercion2.lean b/tests/lean/coercion2.lean index 6d7733106..44e2443d6 100644 --- a/tests/lean/coercion2.lean +++ b/tests/lean/coercion2.lean @@ -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 diff --git a/tests/lean/compact_def.lean b/tests/lean/compact_def.lean index 5d9b54d82..2a94cc5f1 100644 --- a/tests/lean/compact_def.lean +++ b/tests/lean/compact_def.lean @@ -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 \ No newline at end of file +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 \ No newline at end of file diff --git a/tests/lean/compact_def.lean.expected.out b/tests/lean/compact_def.lean.expected.out index de03baca4..9961bf758 100644 --- a/tests/lean/compact_def.lean.expected.out +++ b/tests/lean/compact_def.lean.expected.out @@ -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) diff --git a/tests/lean/config.lean b/tests/lean/config.lean index a24d59dc4..9386ad12f 100644 --- a/tests/lean/config.lean +++ b/tests/lean/config.lean @@ -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 diff --git a/tests/lean/conv.lean b/tests/lean/conv.lean index 173d44be1..3438cf8db 100644 --- a/tests/lean/conv.lean +++ b/tests/lean/conv.lean @@ -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 \ No newline at end of file +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 \ No newline at end of file diff --git a/tests/lean/discharge.lean b/tests/lean/discharge.lean index e877dd9ab..a5d17eab4 100644 --- a/tests/lean/discharge.lean +++ b/tests/lean/discharge.lean @@ -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. diff --git a/tests/lean/disj1.lean b/tests/lean/disj1.lean index 0b46b503e..13102b87e 100644 --- a/tests/lean/disj1.lean +++ b/tests/lean/disj1.lean @@ -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. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/disj1.lean.expected.out b/tests/lean/disj1.lean.expected.out index 2341edfa6..1aefd3eed 100644 --- a/tests/lean/disj1.lean.expected.out +++ b/tests/lean/disj1.lean.expected.out @@ -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)) diff --git a/tests/lean/disjcases.lean b/tests/lean/disjcases.lean index fb205065f..36ed7d8b4 100644 --- a/tests/lean/disjcases.lean +++ b/tests/lean/disjcases.lean @@ -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. diff --git a/tests/lean/elab1.lean b/tests/lean/elab1.lean index 69952a74e..24e6ed3e8 100644 --- a/tests/lean/elab1.lean +++ b/tests/lean/elab1.lean @@ -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))) diff --git a/tests/lean/elab2.lean b/tests/lean/elab2.lean index 84fc773f3..0b4485a31 100644 --- a/tests/lean/elab2.lean +++ b/tests/lean/elab2.lean @@ -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 diff --git a/tests/lean/elab2.lean.expected.out b/tests/lean/elab2.lean.expected.out index b2e168736..82468d123 100644 --- a/tests/lean/elab2.lean.expected.out +++ b/tests/lean/elab2.lean.expected.out @@ -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 diff --git a/tests/lean/elab3.lean b/tests/lean/elab3.lean index c07795713..9eb0cd8e5 100644 --- a/tests/lean/elab3.lean +++ b/tests/lean/elab3.lean @@ -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. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/elab3.lean.expected.out b/tests/lean/elab3.lean.expected.out index 053d4b78a..d75e6f4e3 100644 --- a/tests/lean/elab3.lean.expected.out +++ b/tests/lean/elab3.lean.expected.out @@ -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 diff --git a/tests/lean/elab4.lean b/tests/lean/elab4.lean index 1429b12ff..4af3e58b0 100644 --- a/tests/lean/elab4.lean +++ b/tests/lean/elab4.lean @@ -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. diff --git a/tests/lean/elab4.lean.expected.out b/tests/lean/elab4.lean.expected.out index ffe4a0e8b..9901889a8 100644 --- a/tests/lean/elab4.lean.expected.out +++ b/tests/lean/elab4.lean.expected.out @@ -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 diff --git a/tests/lean/elab5.lean b/tests/lean/elab5.lean index f699278a4..1af323df9 100644 --- a/tests/lean/elab5.lean +++ b/tests/lean/elab5.lean @@ -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. diff --git a/tests/lean/elab5.lean.expected.out b/tests/lean/elab5.lean.expected.out index ffe4a0e8b..9901889a8 100644 --- a/tests/lean/elab5.lean.expected.out +++ b/tests/lean/elab5.lean.expected.out @@ -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 diff --git a/tests/lean/elab7.lean b/tests/lean/elab7.lean index e77a4029d..74d48d06a 100644 --- a/tests/lean/elab7.lean +++ b/tests/lean/elab7.lean @@ -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 \ No newline at end of file +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 \ No newline at end of file diff --git a/tests/lean/elab7.lean.expected.out b/tests/lean/elab7.lean.expected.out index bb2104a2f..4c4ed2ec6 100644 --- a/tests/lean/elab7.lean.expected.out +++ b/tests/lean/elab7.lean.expected.out @@ -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) diff --git a/tests/lean/eq1.lean b/tests/lean/eq1.lean index a0521e529..cd740150a 100644 --- a/tests/lean/eq1.lean +++ b/tests/lean/eq1.lean @@ -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 diff --git a/tests/lean/eq2.lean b/tests/lean/eq2.lean index 9daddfae8..87527eff2 100644 --- a/tests/lean/eq2.lean +++ b/tests/lean/eq2.lean @@ -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. diff --git a/tests/lean/eq3.lean b/tests/lean/eq3.lean index b47afe423..148fc260d 100644 --- a/tests/lean/eq3.lean +++ b/tests/lean/eq3.lean @@ -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 diff --git a/tests/lean/errmsg1.lean b/tests/lean/errmsg1.lean index 4a71e7420..294ba9370 100644 --- a/tests/lean/errmsg1.lean +++ b/tests/lean/errmsg1.lean @@ -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. diff --git a/tests/lean/ex2.lean b/tests/lean/ex2.lean index 8797dcebd..c6c71c8ad 100644 --- a/tests/lean/ex2.lean +++ b/tests/lean/ex2.lean @@ -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 diff --git a/tests/lean/ex2.lean.expected.out b/tests/lean/ex2.lean.expected.out index f0f72d9c4..672580918 100644 --- a/tests/lean/ex2.lean.expected.out +++ b/tests/lean/ex2.lean.expected.out @@ -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 diff --git a/tests/lean/ex3.lean b/tests/lean/ex3.lean index 6304f745b..3670ddb92 100644 --- a/tests/lean/ex3.lean +++ b/tests/lean/ex3.lean @@ -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 \ No newline at end of file +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 \ No newline at end of file diff --git a/tests/lean/exists1.lean b/tests/lean/exists1.lean index 1118bc6b4..11f4c165f 100644 --- a/tests/lean/exists1.lean +++ b/tests/lean/exists1.lean @@ -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. \ No newline at end of file +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. \ No newline at end of file diff --git a/tests/lean/exists1.lean.expected.out b/tests/lean/exists1.lean.expected.out index bfe6770bb..453b01627 100644 --- a/tests/lean/exists1.lean.expected.out +++ b/tests/lean/exists1.lean.expected.out @@ -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 diff --git a/tests/lean/exists2.lean b/tests/lean/exists2.lean index 2b30c73da..a1106e415 100644 --- a/tests/lean/exists2.lean +++ b/tests/lean/exists2.lean @@ -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. \ No newline at end of file +print environment 8. \ No newline at end of file diff --git a/tests/lean/exists2.lean.expected.out b/tests/lean/exists2.lean.expected.out index 5a106c7d7..f9024920b 100644 --- a/tests/lean/exists2.lean.expected.out +++ b/tests/lean/exists2.lean.expected.out @@ -16,11 +16,11 @@ Proved: T6 Proved: T7 Proved: T8 -Theorem T1 : ∃ x y : ℤ, P (f y x) (f y x) := ExistsIntro (g a) (ExistsIntro a H1) -Theorem T2 : ∃ x : ℤ, P (f x (g x)) (f x (g x)) := ExistsIntro a H1 -Theorem T3 : ∃ x : ℤ, P (f x x) (f x x) := ExistsIntro (g a) H2 -Theorem T4 : ∃ x : ℤ, P (f (g a) x) (f x x) := ExistsIntro (g a) H2 -Theorem T5 : ∃ x : ℤ, P x x := ExistsIntro (f (g a) (g a)) H2 -Theorem T6 : ∃ x y : ℤ, P x y := ExistsIntro (f (g a) (g a)) (ExistsIntro (g a) H3) -Theorem T7 : ∃ x : ℤ, P (f x x) x := ExistsIntro (g a) H3 -Theorem T8 : ∃ x y : ℤ, P (f x x) y := ExistsIntro (g a) (ExistsIntro (g a) H3) +theorem T1 : ∃ x y : ℤ, P (f y x) (f y x) := ExistsIntro (g a) (ExistsIntro a H1) +theorem T2 : ∃ x : ℤ, P (f x (g x)) (f x (g x)) := ExistsIntro a H1 +theorem T3 : ∃ x : ℤ, P (f x x) (f x x) := ExistsIntro (g a) H2 +theorem T4 : ∃ x : ℤ, P (f (g a) x) (f x x) := ExistsIntro (g a) H2 +theorem T5 : ∃ x : ℤ, P x x := ExistsIntro (f (g a) (g a)) H2 +theorem T6 : ∃ x y : ℤ, P x y := ExistsIntro (f (g a) (g a)) (ExistsIntro (g a) H3) +theorem T7 : ∃ x : ℤ, P (f x x) x := ExistsIntro (g a) H3 +theorem T8 : ∃ x y : ℤ, P (f x x) y := ExistsIntro (g a) (ExistsIntro (g a) H3) diff --git a/tests/lean/exists3.lean b/tests/lean/exists3.lean index 5611d3258..5d83250d1 100644 --- a/tests/lean/exists3.lean +++ b/tests/lean/exists3.lean @@ -1,16 +1,16 @@ -Import Int. -Variable P : Int -> Int -> Bool +import Int. +variable P : Int -> Int -> Bool -SetOpaque exists false. +setopaque exists false. -Theorem T1 (R1 : not (exists x y, P x y)) : forall x y, not (P x y) := +theorem T1 (R1 : not (exists x y, P x y)) : forall x y, not (P x y) := ForallIntro (fun a, ForallIntro (fun b, ForallElim (DoubleNegElim (ForallElim (DoubleNegElim R1) a)) b)) -Axiom Ax : forall x, exists y, P x y +axiom Ax : forall x, exists y, P x y -Theorem T2 : exists x y, P x y := +theorem T2 : exists x y, P x y := Refute (fun R : not (exists x y, P x y), let L1 : forall x y, not (P x y) := ForallIntro (fun a, ForallIntro (fun b, @@ -19,7 +19,7 @@ Theorem T2 : exists x y, P x y := in ExistsElim L2 (fun (w : Int) (H : P 0 w), Absurd H (ForallElim (ForallElim L1 0) w))). -Theorem T3 (A : (Type U)) (P : A -> A -> Bool) (a : A) (H1 : forall x, exists y, P x y) : exists x y, P x y := +theorem T3 (A : (Type U)) (P : A -> A -> Bool) (a : A) (H1 : forall x, exists y, P x y) : exists x y, P x y := Refute (fun R : not (exists x y, P x y), let L1 : forall x y, not (P x y) := ForallIntro (fun a, ForallIntro (fun b, diff --git a/tests/lean/exists4.lean b/tests/lean/exists4.lean index 97e2fee94..3a1d5934e 100644 --- a/tests/lean/exists4.lean +++ b/tests/lean/exists4.lean @@ -1,16 +1,16 @@ -Variable N : Type -Variables a b c : N -Variables P : N -> N -> N -> Bool -Axiom H3 : P a b c +variable N : Type +variables a b c : N +variables P : N -> N -> N -> Bool +axiom H3 : P a b c -Theorem T1 : exists x y z : N, P x y z := @ExistsIntro N (fun x : N, exists y z : N, P x y z) a +theorem T1 : exists x y z : N, P x y z := @ExistsIntro N (fun x : N, exists y z : N, P x y z) a (@ExistsIntro N _ b (@ExistsIntro N (fun z : N, P a b z) c H3)) -Theorem T2 : exists x y z : N, P x y z := ExistsIntro a (ExistsIntro b (ExistsIntro c H3)) +theorem T2 : exists x y z : N, P x y z := ExistsIntro a (ExistsIntro b (ExistsIntro c H3)) -Theorem T3 : exists x y z : N, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H3)) +theorem T3 : exists x y z : N, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H3)) -Theorem T4 (H : P a a b) : exists x y z, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H)) +theorem T4 (H : P a a b) : exists x y z, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H)) -print Environment 4 \ No newline at end of file +print environment 4 \ No newline at end of file diff --git a/tests/lean/exists4.lean.expected.out b/tests/lean/exists4.lean.expected.out index 582b65c90..1bb871e20 100644 --- a/tests/lean/exists4.lean.expected.out +++ b/tests/lean/exists4.lean.expected.out @@ -10,12 +10,12 @@ Proved: T2 Proved: T3 Proved: T4 -Theorem T1 : ∃ x y z : N, P x y z := +theorem T1 : ∃ x y z : N, P x y z := @ExistsIntro N (λ x : N, ∃ y z : N, P x y z) a (@ExistsIntro N (λ y : N, ∃ z : N, P a y z) b (@ExistsIntro N (λ z : N, P a b z) c H3)) -Theorem T2 : ∃ x y z : N, P x y z := ExistsIntro a (ExistsIntro b (ExistsIntro c H3)) -Theorem T3 : ∃ x y z : N, P x y z := ExistsIntro a (ExistsIntro b (ExistsIntro c H3)) -Theorem T4 (H : P a a b) : ∃ x y z : N, P x y z := ExistsIntro a (ExistsIntro a (ExistsIntro b H)) +theorem T2 : ∃ x y z : N, P x y z := ExistsIntro a (ExistsIntro b (ExistsIntro c H3)) +theorem T3 : ∃ x y z : N, P x y z := ExistsIntro a (ExistsIntro b (ExistsIntro c H3)) +theorem T4 (H : P a a b) : ∃ x y z : N, P x y z := ExistsIntro a (ExistsIntro a (ExistsIntro b H)) diff --git a/tests/lean/exists5.lean b/tests/lean/exists5.lean index f3dc7203d..fcfe6454f 100644 --- a/tests/lean/exists5.lean +++ b/tests/lean/exists5.lean @@ -1,7 +1,7 @@ -Variable N : Type -Variables a b c : N -Variables P : N -> N -> N -> Bool +variable N : Type +variables a b c : N +variables P : N -> N -> N -> Bool -Theorem T1 (f : N -> N) (H : P (f a) b (f (f c))) : exists x y z, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H)) +theorem T1 (f : N -> N) (H : P (f a) b (f (f c))) : exists x y z, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H)) -print Environment 1. +print environment 1. diff --git a/tests/lean/exists5.lean.expected.out b/tests/lean/exists5.lean.expected.out index 8f7186e28..675ac2754 100644 --- a/tests/lean/exists5.lean.expected.out +++ b/tests/lean/exists5.lean.expected.out @@ -6,5 +6,5 @@ Assumed: c Assumed: P Proved: T1 -Theorem T1 (f : N → N) (H : P (f a) b (f (f c))) : ∃ x y z : N, P x y z := +theorem T1 (f : N → N) (H : P (f a) b (f (f c))) : ∃ x y z : N, P x y z := ExistsIntro (f a) (ExistsIntro b (ExistsIntro (f (f c)) H)) diff --git a/tests/lean/exists6.lean b/tests/lean/exists6.lean index e99c74616..b04aa3b73 100644 --- a/tests/lean/exists6.lean +++ b/tests/lean/exists6.lean @@ -1,8 +1,8 @@ -Import Int. -Variable P : Int -> Int -> Int -> Bool -Axiom Ax1 : exists x y z, P x y z -Axiom Ax2 : forall x y z, not (P x y z) -Theorem T : false := +import Int. +variable P : Int -> Int -> Int -> Bool +axiom Ax1 : exists x y z, P x y z +axiom Ax2 : forall x y z, not (P x y z) +theorem T : false := ExistsElim Ax1 (fun a H1, ExistsElim H1 (fun b H2, ExistsElim H2 (fun (c : Int) (H3 : P a b c), let notH3 : not (P a b c) := ForallElim (ForallElim (ForallElim Ax2 a) b) c in Absurd H3 notH3))) diff --git a/tests/lean/exists7.lean b/tests/lean/exists7.lean index 59bec1bcb..491dc4fe4 100644 --- a/tests/lean/exists7.lean +++ b/tests/lean/exists7.lean @@ -1,12 +1,12 @@ -SetOption pp::colors false -Variable N : Type -Variables a b c : N -Variables P : N -> N -> N -> Bool +setoption pp::colors false +variable N : Type +variables a b c : N +variables P : N -> N -> N -> Bool -SetOpaque forall false. -SetOpaque exists false. -SetOpaque not false. +setopaque forall false. +setopaque exists false. +setopaque not false. -Theorem T1 (f : N -> N) (H : P (f a) b (f (f c))) : exists x y z, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H)) +theorem T1 (f : N -> N) (H : P (f a) b (f (f c))) : exists x y z, P x y z := ExistsIntro _ (ExistsIntro _ (ExistsIntro _ H)) -print Environment 1. +print environment 1. diff --git a/tests/lean/exists7.lean.expected.out b/tests/lean/exists7.lean.expected.out index 4c8c1fb8d..620a488be 100644 --- a/tests/lean/exists7.lean.expected.out +++ b/tests/lean/exists7.lean.expected.out @@ -7,5 +7,5 @@ Assumed: c Assumed: P Proved: T1 -Theorem T1 (f : N → N) (H : P (f a) b (f (f c))) : ∃ x y z : N, P x y z := +theorem T1 (f : N → N) (H : P (f a) b (f (f c))) : ∃ x y z : N, P x y z := ExistsIntro (f a) (ExistsIntro b (ExistsIntro (f (f c)) H)) diff --git a/tests/lean/exists8.lean b/tests/lean/exists8.lean index f6369190c..587132870 100644 --- a/tests/lean/exists8.lean +++ b/tests/lean/exists8.lean @@ -1,14 +1,14 @@ -Import Int. -Variable P : Int -> Int -> Bool +import Int. +variable P : Int -> Int -> Bool -Theorem T1 (R1 : not (exists x y, P x y)) : forall x y, not (P x y) := +theorem T1 (R1 : not (exists x y, P x y)) : forall x y, not (P x y) := ForallIntro (fun a, ForallIntro (fun b, ForallElim (NotExistsElim (ForallElim (NotExistsElim R1) a)) b)) -Axiom Ax : forall x, exists y, P x y +axiom Ax : forall x, exists y, P x y -Theorem T2 : exists x y, P x y := +theorem T2 : exists x y, P x y := Refute (fun R : not (exists x y, P x y), let L1 : forall x y, not (P x y) := ForallIntro (fun a, ForallIntro (fun b, @@ -17,7 +17,7 @@ Theorem T2 : exists x y, P x y := in ExistsElim L2 (fun (w : Int) (H : P 0 w), Absurd H (ForallElim (ForallElim L1 0) w))). -Theorem T3 (A : (Type U)) (P : A -> A -> Bool) (a : A) (H1 : forall x, exists y, P x y) : exists x y, P x y := +theorem T3 (A : (Type U)) (P : A -> A -> Bool) (a : A) (H1 : forall x, exists y, P x y) : exists x y, P x y := Refute (fun R : not (exists x y, P x y), let L1 : forall x y, not (P x y) := ForallIntro (fun a, ForallIntro (fun b, diff --git a/tests/lean/exit.lean b/tests/lean/exit.lean index 90e050aa7..553d6d277 100644 --- a/tests/lean/exit.lean +++ b/tests/lean/exit.lean @@ -1,2 +1,2 @@ -Exit +exit print "FAILED" diff --git a/tests/lean/explicit.lean b/tests/lean/explicit.lean index f6ad3cd71..6f79e8f94 100644 --- a/tests/lean/explicit.lean +++ b/tests/lean/explicit.lean @@ -1,9 +1,9 @@ -Import Int. -Variable f {A : Type} : A -> A -> A -Variable module::g {A : Type} : A -> A -> A -Check @f -Check module::@g -Variable h::1 {A B : Type} : A -> B -> A -Check h::1::explicit -Variable @h : Int -> Int -Variable h {A B : Type} : A -> B -> A \ No newline at end of file +import Int. +variable f {A : Type} : A -> A -> A +variable module::g {A : Type} : A -> A -> A +check @f +check module::@g +variable h::1 {A B : Type} : A -> B -> A +check h::1::explicit +variable @h : Int -> Int +variable h {A B : Type} : A -> B -> A \ No newline at end of file diff --git a/tests/lean/find.lean b/tests/lean/find.lean index 76f13e7cf..057dd4b3d 100644 --- a/tests/lean/find.lean +++ b/tests/lean/find.lean @@ -1,4 +1,4 @@ -(* import("find.lua") *) -Find "^.ongr" -Find "foo" -Find "(ab" \ No newline at end of file +import find. +find "^.ongr" +find "foo" +find "(ab" \ No newline at end of file diff --git a/tests/lean/find.lean.expected.out b/tests/lean/find.lean.expected.out index ed935d97e..9ba9e80c1 100644 --- a/tests/lean/find.lean.expected.out +++ b/tests/lean/find.lean.expected.out @@ -1,7 +1,8 @@ Set: pp::colors Set: pp::unicode -Theorem Congr1 {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} (a : A) (H : f == g) : f a == g a -Theorem Congr2 {A : TypeU} {B : A → TypeU} {a b : A} (f : Π x : A, B x) (H : a == b) : f a == f b -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 + Imported 'find' +theorem Congr1 {A : TypeU} {B : A → TypeU} {f g : Π x : A, B x} (a : A) (H : f == g) : f a == g a +theorem Congr2 {A : TypeU} {B : A → TypeU} {a b : A} (f : Π x : A, B x) (H : a == b) : f a == f b +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 Error (line: 3, pos: 0) executing external script (/home/leo/projects/lean/build/debug/shell/find.lua:24), no object name in the environment matches the regular expression 'foo' -Error (line: 4, pos: 0) executing external script (/home/leo/projects/lean/build/debug/shell/find.lua:18), unfinished capture +Error (line: 4, pos: 0) executing external script (/home/leo/projects/lean/build/debug/shell/find.lua:24), no object name in the environment matches the regular expression '(ab' diff --git a/tests/lean/forall1.lean b/tests/lean/forall1.lean index e3b19d398..6a07ee752 100644 --- a/tests/lean/forall1.lean +++ b/tests/lean/forall1.lean @@ -1,4 +1,4 @@ -Import Int. -Variable P : Int -> Bool -Axiom Ax (x : Int) : P x -Check ForallIntro Ax \ No newline at end of file +import Int. +variable P : Int -> Bool +axiom Ax (x : Int) : P x +check ForallIntro Ax \ No newline at end of file diff --git a/tests/lean/ho.lean b/tests/lean/ho.lean index 79e340325..368987e8c 100644 --- a/tests/lean/ho.lean +++ b/tests/lean/ho.lean @@ -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 := Subst 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 := Subst H2 H1 +print environment 2 diff --git a/tests/lean/ho.lean.expected.out b/tests/lean/ho.lean.expected.out index 6c7e0c05d..653a0d365 100644 --- a/tests/lean/ho.lean.expected.out +++ b/tests/lean/ho.lean.expected.out @@ -7,5 +7,5 @@ Assumed: H1 Assumed: H2 Proved: T1 -Axiom H2 : g a > 0 -Theorem T1 : g b > 0 := Subst H2 H1 +axiom H2 : g a > 0 +theorem T1 : g b > 0 := Subst H2 H1 diff --git a/tests/lean/implicit1.lean b/tests/lean/implicit1.lean index 91943dde5..4ae6c4b18 100644 --- a/tests/lean/implicit1.lean +++ b/tests/lean/implicit1.lean @@ -1,12 +1,12 @@ -Import Int. -Import Real. -Variable f : Int -> Int -> Int +import Int. +import Real. +variable f : Int -> Int -> Int print forall a, f a a > 0 print forall a b, f a b > 0 -Variable g : Int -> Real -> Int +variable g : Int -> Real -> Int print forall a b, g a b > 0 print forall a b, g a (f a b) > 0 -SetOption pp::coercion true +setoption pp::coercion true print forall a b, g a (f a b) > 0 print fun a, a + 1 print fun a b, a + b @@ -15,6 +15,6 @@ print fun (a b) (c : Int), a + c + b -- The current elaborator resolves overloads before solving the implicit argument constraints. -- So, it does not have enough information for deciding which overload to use. print (fun a b, a + b) 10 20. -Variable x : Int +variable x : Int -- The following one works because the type of x is used to decide which + should be used print fun a b, a + x + b \ No newline at end of file diff --git a/tests/lean/implicit2.lean b/tests/lean/implicit2.lean index 904c3ea98..2cc06e19d 100644 --- a/tests/lean/implicit2.lean +++ b/tests/lean/implicit2.lean @@ -1,13 +1,13 @@ -Import Real. -Variable f {A : Type} (x y : A) : A -Check f 10 20 -Check f 10 -Check @f -Variable g {A : Type} (x1 x2 : A) {B : Type} (y1 y2 : B) : B -Check g 10 20 true -Check let r : Real -> Real -> Real := g 10 20 +import Real. +variable f {A : Type} (x y : A) : A +check f 10 20 +check f 10 +check @f +variable g {A : Type} (x1 x2 : A) {B : Type} (y1 y2 : B) : B +check g 10 20 true +check let r : Real -> Real -> Real := g 10 20 in r -Check g 10 -SetOption pp::implicit true -Check let r : Real -> Real -> Real := g 10 20 +check g 10 +setoption pp::implicit true +check let r : Real -> Real -> Real := g 10 20 in r diff --git a/tests/lean/implicit3.lean b/tests/lean/implicit3.lean index 99af0e59f..636b8caee 100644 --- a/tests/lean/implicit3.lean +++ b/tests/lean/implicit3.lean @@ -1,11 +1,11 @@ -Import Int. +import Int. print 10 = 20 -Variable f : Int -> Int -> Int -Variable g : Int -> Int -> Int -> Int -Notation 10 _ ++ _ : f -Notation 10 _ ++ _ : g -SetOption pp::implicit true -SetOption pp::notation false +variable f : Int -> Int -> Int +variable g : Int -> Int -> Int -> Int +notation 10 _ ++ _ : f +notation 10 _ ++ _ : g +setoption pp::implicit true +setoption pp::notation false print (10 ++ 20) print (10 ++ 20) 10 diff --git a/tests/lean/implicit4.lean b/tests/lean/implicit4.lean index 4d1de7e48..a78b27989 100644 --- a/tests/lean/implicit4.lean +++ b/tests/lean/implicit4.lean @@ -1,9 +1,9 @@ -Import Int. -Variable f {A : Type} (a1 a2 : A) {B : Type} (b1 b2 : B) : A -Variable g {A1 A2 : Type} (a1 : A1) (a2 : A2) {B : Type} (b : B) : A1 -Variable p (a1 a2 : Int) {B : Type} (b1 b2 b3 : B) : B -Variable h {A1 A2 : Type} (a1 : A1) (a2 : A2) (a3 : A2) : A1 -Infix ++ : f -Infix ++ : g -Infix ++ : p -Infix ++ : h \ No newline at end of file +import Int. +variable f {A : Type} (a1 a2 : A) {B : Type} (b1 b2 : B) : A +variable g {A1 A2 : Type} (a1 : A1) (a2 : A2) {B : Type} (b : B) : A1 +variable p (a1 a2 : Int) {B : Type} (b1 b2 b3 : B) : B +variable h {A1 A2 : Type} (a1 : A1) (a2 : A2) (a3 : A2) : A1 +infix ++ : f +infix ++ : g +infix ++ : p +infix ++ : h \ No newline at end of file diff --git a/tests/lean/implicit4.lean.expected.out b/tests/lean/implicit4.lean.expected.out index cd349a46f..180f0297d 100644 --- a/tests/lean/implicit4.lean.expected.out +++ b/tests/lean/implicit4.lean.expected.out @@ -6,7 +6,7 @@ Assumed: p Assumed: h The denotation(s) for the existing notation: - Infix ++ + infix ++ have been replaced with the new denotation: @h because they conflict on how implicit arguments are used. diff --git a/tests/lean/implicit5.lean b/tests/lean/implicit5.lean index cb48594e4..0a47e5ead 100644 --- a/tests/lean/implicit5.lean +++ b/tests/lean/implicit5.lean @@ -1,12 +1,12 @@ -Import Int. -Import Real. -Variable f {A : Type} (a1 a2 : A) : A -Variable g : Int -> Int -> Int -Variable h : Int -> Int -> Real -> Int -Variable p {A B : Type} (a1 a2 : A) (b : B) : A -Infix ++ : f -Infix ++ : g -Infix ++ : h -Infix ++ : p -Variable p2 {A B : Type} (a1 a2 : A) (b : B) {C : Type} (c : C) : A -Infix ++ : p2 +import Int. +import Real. +variable f {A : Type} (a1 a2 : A) : A +variable g : Int -> Int -> Int +variable h : Int -> Int -> Real -> Int +variable p {A B : Type} (a1 a2 : A) (b : B) : A +infix ++ : f +infix ++ : g +infix ++ : h +infix ++ : p +variable p2 {A B : Type} (a1 a2 : A) (b : B) {C : Type} (c : C) : A +infix ++ : p2 diff --git a/tests/lean/implicit5.lean.expected.out b/tests/lean/implicit5.lean.expected.out index cfc4e7755..a4f4f94b8 100644 --- a/tests/lean/implicit5.lean.expected.out +++ b/tests/lean/implicit5.lean.expected.out @@ -8,7 +8,7 @@ Assumed: p Assumed: p2 The denotation(s) for the existing notation: - Infix ++ + infix ++ have been replaced with the new denotation: @p2 because they conflict on how implicit arguments are used. diff --git a/tests/lean/implicit6.lean b/tests/lean/implicit6.lean index 03580073e..6f760811e 100644 --- a/tests/lean/implicit6.lean +++ b/tests/lean/implicit6.lean @@ -1,11 +1,11 @@ -Import Int. -Variable f {A : Type} : A -> A -> A -Infixl 65 + : f +import Int. +variable f {A : Type} : A -> A -> A +infixl 65 + : f print true + false print 10 + 20 print 10 + (- 20) -SetOption pp::notation false -SetOption pp::coercion true +setoption pp::notation false +setoption pp::coercion true print true + false print 10 + 20 print 10 + (- 20) diff --git a/tests/lean/implicit7.lean b/tests/lean/implicit7.lean index 9e2024f63..7f23fb805 100644 --- a/tests/lean/implicit7.lean +++ b/tests/lean/implicit7.lean @@ -1,12 +1,12 @@ -Variable f {A : Type} (a : A) {B : Type} : A -> B -> A -Variable g {A B : Type} (a : A) {C : Type} : B -> C -> C -Notation 100 _ ; _ ; _ : f -Notation 100 _ ; _ ; _ : g -Check 10 ; true ; false -Check 10 ; 10 ; true -SetOption pp::notation false -Check 10 ; true ; false -Check 10 ; 10 ; true -SetOption pp::implicit true -Check 10 ; true ; false -Check 10 ; 10 ; true +variable f {A : Type} (a : A) {B : Type} : A -> B -> A +variable g {A B : Type} (a : A) {C : Type} : B -> C -> C +notation 100 _ ; _ ; _ : f +notation 100 _ ; _ ; _ : g +check 10 ; true ; false +check 10 ; 10 ; true +setoption pp::notation false +check 10 ; true ; false +check 10 ; 10 ; true +setoption pp::implicit true +check 10 ; true ; false +check 10 ; 10 ; true diff --git a/tests/lean/interactive/config.lean b/tests/lean/interactive/config.lean index a24d59dc4..9386ad12f 100644 --- a/tests/lean/interactive/config.lean +++ b/tests/lean/interactive/config.lean @@ -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 diff --git a/tests/lean/interactive/elab6.lean b/tests/lean/interactive/elab6.lean index ba67ad981..7db433aa8 100644 --- a/tests/lean/interactive/elab6.lean +++ b/tests/lean/interactive/elab6.lean @@ -5,5 +5,5 @@ local env = get_environment() env:set_opaque("forall", false) *) -Theorem ForallIntro2 (A : (Type U)) (P : A -> Bool) (H : Pi x, P x) : forall x, P x := +theorem ForallIntro2 (A : (Type U)) (P : A -> Bool) (H : Pi x, P x) : forall x, P x := Abst (fun x, EqTIntro (H x)) diff --git a/tests/lean/interactive/t1.lean b/tests/lean/interactive/t1.lean index acfce3f3c..66e578907 100644 --- a/tests/lean/interactive/t1.lean +++ b/tests/lean/interactive/t1.lean @@ -1,4 +1,4 @@ -Theorem T2 (a b : Bool) : a => b => a /\ b. +theorem T2 (a b : Bool) : a => b => a /\ b. done. done. (* imp_tac() *). diff --git a/tests/lean/interactive/t10.lean b/tests/lean/interactive/t10.lean index c4c7014cd..c2c1d340b 100644 --- a/tests/lean/interactive/t10.lean +++ b/tests/lean/interactive/t10.lean @@ -2,7 +2,7 @@ simple_tac = Repeat(OrElse(conj_hyp_tac(), conj_tac(), assumption_tac())) *) -Theorem T2 (A B : Bool) : A /\ B => B /\ A := +theorem T2 (A B : Bool) : A /\ B => B /\ A := Discharge (fun H : A /\ B, let H1 : A := _, H2 : B := _, diff --git a/tests/lean/interactive/t11.lean b/tests/lean/interactive/t11.lean index 7a19b0014..f85678373 100644 --- a/tests/lean/interactive/t11.lean +++ b/tests/lean/interactive/t11.lean @@ -2,7 +2,7 @@ auto = Repeat(OrElse(conj_hyp_tac(), conj_tac(), assumption_tac())) *) -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 := _, lemma2 : B := _, diff --git a/tests/lean/interactive/t12.lean b/tests/lean/interactive/t12.lean index 0c01833c1..af5fc8b62 100644 --- a/tests/lean/interactive/t12.lean +++ b/tests/lean/interactive/t12.lean @@ -4,15 +4,15 @@ import("tactic.lua") auto = Repeat(OrElse(assumption_tac(), conj_tac(), conj_hyp_tac())) *) -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 -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 := _, lemma2 : B := _ @@ -21,7 +21,7 @@ Theorem T2 (A B : Bool) : A /\ B -> B /\ A := auto. done. auto. done. -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 := _, lemma2 : B := _ @@ -30,7 +30,7 @@ Theorem T3 (A B : Bool) : A /\ B -> B /\ A := conj_hyp. exact. done. apply Conj. exact. done. -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 := _, lemma2 : B := _ diff --git a/tests/lean/interactive/t12.lean.expected.out b/tests/lean/interactive/t12.lean.expected.out index a2f46e3b3..035beb6eb 100644 --- a/tests/lean/interactive/t12.lean.expected.out +++ b/tests/lean/interactive/t12.lean.expected.out @@ -1,7 +1,7 @@ # Set: pp::colors Set: pp::unicode Proved: T1 -Theorem T1 (A B : Bool) (assumption : A ∧ B) : B ∧ A := +theorem T1 (A B : Bool) (assumption : A ∧ B) : B ∧ A := let lemma1 : A := Conjunct1 assumption, lemma2 : B := Conjunct2 assumption in Conj lemma2 lemma1 # Proof state: A : Bool, B : Bool, assumption : A ∧ B ⊢ A diff --git a/tests/lean/interactive/t13.lean b/tests/lean/interactive/t13.lean index 5b14024b6..be5644f82 100644 --- a/tests/lean/interactive/t13.lean +++ b/tests/lean/interactive/t13.lean @@ -3,10 +3,10 @@ auto = Repeat(OrElse(assumption_tac(), conj_tac(), conj_hyp_tac())) *) -Theorem T1 (A B : Bool) : A /\ B -> B /\ A := +theorem T1 (A B : Bool) : A /\ B -> B /\ A := fun assumption : A /\ B, let lemma1 := (have A by auto), lemma2 := (have B by auto) in (have B /\ A by auto) -print Environment 1. +print environment 1. diff --git a/tests/lean/interactive/t13.lean.expected.out b/tests/lean/interactive/t13.lean.expected.out index 8d1cda17c..c34bb68de 100644 --- a/tests/lean/interactive/t13.lean.expected.out +++ b/tests/lean/interactive/t13.lean.expected.out @@ -1,6 +1,6 @@ # Set: pp::colors Set: pp::unicode Proved: T1 -Theorem T1 (A B : Bool) (assumption : A ∧ B) : B ∧ A := +theorem T1 (A B : Bool) (assumption : A ∧ B) : B ∧ A := let lemma1 := Conjunct1 assumption, lemma2 := Conjunct2 assumption in Conj lemma2 lemma1 # \ No newline at end of file diff --git a/tests/lean/interactive/t2.lean b/tests/lean/interactive/t2.lean index f07ff5723..5fad68a4a 100644 --- a/tests/lean/interactive/t2.lean +++ b/tests/lean/interactive/t2.lean @@ -1,9 +1,9 @@ -Theorem T2 (a b : Bool) : a => b => a /\ b. +theorem T2 (a b : Bool) : a => b => a /\ b. (* imp_tac() *) (* imp_tac2() *) foo. (* imp_tac() *) abort. -Variables a b : Bool. -print Environment 2. +variables a b : Bool. +print environment 2. diff --git a/tests/lean/interactive/t2.lean.expected.out b/tests/lean/interactive/t2.lean.expected.out index bcde59c4b..4c5b21e08 100644 --- a/tests/lean/interactive/t2.lean.expected.out +++ b/tests/lean/interactive/t2.lean.expected.out @@ -8,6 +8,6 @@ Proof state: H::1 : b, H : a, a : Bool, b : Bool ⊢ a ∧ b # Assumed: a Assumed: b -# Variable a : Bool -Variable b : Bool +# variable a : Bool +variable b : Bool # \ No newline at end of file diff --git a/tests/lean/interactive/t3.lean b/tests/lean/interactive/t3.lean index 2e61398c6..03e933d5a 100644 --- a/tests/lean/interactive/t3.lean +++ b/tests/lean/interactive/t3.lean @@ -1,6 +1,6 @@ (* import("tactic.lua") *) -Theorem T2 (a b : Bool) : b => a \/ b. +theorem T2 (a b : Bool) : b => a \/ b. (* imp_tac() *). (* disj_tac() *). back. @@ -8,4 +8,4 @@ back. exact. done. -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/interactive/t3.lean.expected.out b/tests/lean/interactive/t3.lean.expected.out index 52d917a1c..0e61a9c31 100644 --- a/tests/lean/interactive/t3.lean.expected.out +++ b/tests/lean/interactive/t3.lean.expected.out @@ -14,5 +14,5 @@ H : b, a : Bool, b : Bool ⊢ b ## Proof state: no goals ## Proved: T2 -# Theorem T2 (a b : Bool) : b ⇒ a ∨ b := Discharge (λ H : b, Disj2 a H) +# theorem T2 (a b : Bool) : b ⇒ a ∨ b := Discharge (λ H : b, Disj2 a H) # \ No newline at end of file diff --git a/tests/lean/interactive/t4.lean b/tests/lean/interactive/t4.lean index 46f967abb..7f73e6127 100644 --- a/tests/lean/interactive/t4.lean +++ b/tests/lean/interactive/t4.lean @@ -1,2 +1,2 @@ -Theorem (a : Bool) : a. +theorem (a : Bool) : a. print "done". \ No newline at end of file diff --git a/tests/lean/interactive/t5.lean b/tests/lean/interactive/t5.lean index 309c224df..f11ffd685 100644 --- a/tests/lean/interactive/t5.lean +++ b/tests/lean/interactive/t5.lean @@ -1,8 +1,8 @@ (* import("tactic.lua") *) -Axiom magic (a : Bool) : a. +axiom magic (a : Bool) : a. -Theorem T (a : Bool) : a. +theorem T (a : Bool) : a. apply magic. done. -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/interactive/t5.lean.expected.out b/tests/lean/interactive/t5.lean.expected.out index 1dd8ef3a0..402f898f5 100644 --- a/tests/lean/interactive/t5.lean.expected.out +++ b/tests/lean/interactive/t5.lean.expected.out @@ -6,5 +6,5 @@ a : Bool ⊢ a ## Proof state: no goals ## Proved: T -# Theorem T (a : Bool) : a := magic a +# theorem T (a : Bool) : a := magic a # \ No newline at end of file diff --git a/tests/lean/interactive/t6.lean b/tests/lean/interactive/t6.lean index 5d4100e9e..ba666b779 100644 --- a/tests/lean/interactive/t6.lean +++ b/tests/lean/interactive/t6.lean @@ -1,5 +1,5 @@ (* import("tactic.lua") *) -Theorem T1 (a b : Bool) : a => b => a /\ b. +theorem T1 (a b : Bool) : a => b => a /\ b. (* imp_tac() *). (* imp_tac() *). apply Conj. diff --git a/tests/lean/interactive/t7.lean b/tests/lean/interactive/t7.lean index 24f138fd2..15849d843 100644 --- a/tests/lean/interactive/t7.lean +++ b/tests/lean/interactive/t7.lean @@ -1,10 +1,10 @@ -Import Int. +import Int. (* import("tactic.lua") *) -Variable q : Int -> Int -> Bool. -Variable p : Int -> Bool. -Axiom Ax (a b : Int) (H : q a b) : p b. -Variable a : Int. -Theorem T (x : Int) : (q a x) => (p x). +variable q : Int -> Int -> Bool. +variable p : Int -> Bool. +axiom Ax (a b : Int) (H : q a b) : p b. +variable a : Int. +theorem T (x : Int) : (q a x) => (p x). (* imp_tac() *). apply (Ax a). exact. diff --git a/tests/lean/interactive/t8.lean b/tests/lean/interactive/t8.lean index c5213063a..c09f14d21 100644 --- a/tests/lean/interactive/t8.lean +++ b/tests/lean/interactive/t8.lean @@ -1,6 +1,6 @@ (* import("tactic.lua") *) -SetOption tactic::proof_state::goal_names true. -Theorem T (a : Bool) : a => a /\ a. +setoption tactic::proof_state::goal_names true. +theorem T (a : Bool) : a => a /\ a. apply Discharge. apply Conj. exact. diff --git a/tests/lean/interactive/t9.lean b/tests/lean/interactive/t9.lean index 8cf93cc86..48fb412e1 100644 --- a/tests/lean/interactive/t9.lean +++ b/tests/lean/interactive/t9.lean @@ -1,5 +1,5 @@ (* import("tactic.lua") *) -Theorem T1 (A B : Bool) : A /\ B => B /\ A := +theorem T1 (A B : Bool) : A /\ B => B /\ A := Discharge (fun H : A /\ B, let main : B /\ A := (let H1 : B := _, @@ -20,7 +20,7 @@ Theorem T1 (A B : Bool) : A /\ B => B /\ A := simple_tac = Repeat(OrElse(conj_hyp_tac(), conj_tac(), assumption_tac())) *) -Theorem T2 (A B : Bool) : A /\ B => B /\ A := +theorem T2 (A B : Bool) : A /\ B => B /\ A := Discharge (fun H : A /\ B, let H1 : A := _, H2 : B := _, diff --git a/tests/lean/let1.lean b/tests/lean/let1.lean index 5e96f2f12..12318bd3c 100644 --- a/tests/lean/let1.lean +++ b/tests/lean/let1.lean @@ -1,10 +1,10 @@ -Import Int. +import Int. print let a : Nat := 10, b : Nat := 20, c : Nat := 30, d : Nat := 10 in a + b + c + d print let a : Nat := 1000000000000000000, b : Nat := 20000000000000000000, c : Nat := 3000000000000000000, d : Nat := 4000000000000000000 in a + b + c + d -Check let a : Nat := 10 in a + 1 -Eval let a : Nat := 20 in a + 10 -Eval let a := 20 in a + 10 -Check let a : Int := 20 in a + 10 -SetOption pp::coercion true +check let a : Nat := 10 in a + 1 +eval let a : Nat := 20 in a + 10 +eval let a := 20 in a + 10 +check let a : Int := 20 in a + 10 +setoption pp::coercion true print let a : Int := 20 in a + 10 diff --git a/tests/lean/let2.lean b/tests/lean/let2.lean index aae077c71..0653180f0 100644 --- a/tests/lean/let2.lean +++ b/tests/lean/let2.lean @@ -1,10 +1,10 @@ -- Annotating lemmas -Theorem simple (p q r : Bool) : (p ⇒ q) ∧ (q ⇒ r) ⇒ p ⇒ r := +theorem simple (p q r : Bool) : (p ⇒ q) ∧ (q ⇒ r) ⇒ p ⇒ r := Discharge (λ H_pq_qr, Discharge (λ H_p, let P_pq : (p ⇒ q) := Conjunct1 H_pq_qr, P_qr : (q ⇒ r) := Conjunct2 H_pq_qr, P_q : q := MP P_pq H_p in MP P_qr P_q)) -print Environment 1 \ No newline at end of file +print environment 1 \ No newline at end of file diff --git a/tests/lean/let2.lean.expected.out b/tests/lean/let2.lean.expected.out index 9fd01d6f6..de9904e32 100644 --- a/tests/lean/let2.lean.expected.out +++ b/tests/lean/let2.lean.expected.out @@ -1,7 +1,7 @@ Set: pp::colors Set: pp::unicode Proved: simple -Theorem simple (p q r : Bool) : (p ⇒ q) ∧ (q ⇒ r) ⇒ p ⇒ r := +theorem simple (p q r : Bool) : (p ⇒ q) ∧ (q ⇒ r) ⇒ p ⇒ r := Discharge (λ H_pq_qr : (p ⇒ q) ∧ (q ⇒ r), Discharge diff --git a/tests/lean/let3.lean b/tests/lean/let3.lean index 938d4b44c..cda1b254b 100644 --- a/tests/lean/let3.lean +++ b/tests/lean/let3.lean @@ -1,9 +1,9 @@ -Import Int. +import Int. -Variable magic : Pi (H : Bool), H +variable magic : Pi (H : Bool), H -SetOption pp::notation false -SetOption pp::coercion true +setoption pp::notation false +setoption pp::coercion true print let a : Int := 1, H : a > 0 := magic (a > 0) in H \ No newline at end of file diff --git a/tests/lean/let4.lean b/tests/lean/let4.lean index 7d203ef2a..74e672913 100644 --- a/tests/lean/let4.lean +++ b/tests/lean/let4.lean @@ -1,14 +1,14 @@ -Import Int. +import Int. print let b := true, a : Int := b in a -Variable vector : Type -> Nat -> Type -Variable const {A : Type} (n : Nat) (a : A) : vector A n +variable vector : Type -> Nat -> Type +variable const {A : Type} (n : Nat) (a : A) : vector A n -Check +check let a := 10, v1 := const a true, v2 := v1 @@ -20,28 +20,28 @@ let a := 10, v2 : vector Bool a := v1 in v2 -Check +check let a := 10, v1 : vector Bool a := const a true, v2 : vector Bool a := v1 in v2 -Check +check let a := 10, v1 : vector Bool a := const a true, v2 : vector Int a := v1 in v2 -Variable foo : (vector Bool 10) -> (vector Int 10) -Coercion foo +variable foo : (vector Bool 10) -> (vector Int 10) +coercion foo -Check +check let a := 10, v1 : vector Bool a := const a true, v2 : vector Int a := v1 in v2 -SetOption pp::coercion true +setoption pp::coercion true print let a := 10, diff --git a/tests/lean/level1.lean b/tests/lean/level1.lean index 7762ae223..257fe6e79 100644 --- a/tests/lean/level1.lean +++ b/tests/lean/level1.lean @@ -1 +1 @@ -Variable T : (Type (max U+1 U+2)). +variable T : (Type (max U+1 U+2)). diff --git a/tests/lean/loop2.lean b/tests/lean/loop2.lean index 96b40ac51..9fc6e08bf 100644 --- a/tests/lean/loop2.lean +++ b/tests/lean/loop2.lean @@ -1,4 +1,4 @@ -Import Int. +import Int. (* function add_paren(code) diff --git a/tests/lean/loop2.lean.expected.out b/tests/lean/loop2.lean.expected.out index 464d6b7af..e1ca8f684 100644 --- a/tests/lean/loop2.lean.expected.out +++ b/tests/lean/loop2.lean.expected.out @@ -1,5 +1,5 @@ Set: pp::colors Set: pp::unicode Imported 'Int' -Variable x : ℤ +variable x : ℤ done diff --git a/tests/lean/lua1.lean b/tests/lean/lua1.lean index 78196067a..b72bb68e3 100644 --- a/tests/lean/lua1.lean +++ b/tests/lean/lua1.lean @@ -1,8 +1,8 @@ -Import Int. -Variable x : Int +import Int. +variable x : Int (* print("hello world from Lua") *) -Variable y : Int +variable y : Int diff --git a/tests/lean/lua10.lean b/tests/lean/lua10.lean index ce433d55f..b19a99bf0 100644 --- a/tests/lean/lua10.lean +++ b/tests/lean/lua10.lean @@ -1,5 +1,5 @@ -Variables x1 x2 x3 : Bool -Definition F : Bool := x1 /\ (x2 \/ x3) +variables x1 x2 x3 : Bool +definition F : Bool := x1 /\ (x2 \/ x3) (* local env = get_environment() diff --git a/tests/lean/lua11.lean b/tests/lean/lua11.lean index c59532d4d..e56e34452 100644 --- a/tests/lean/lua11.lean +++ b/tests/lean/lua11.lean @@ -1,4 +1,4 @@ -Import Int. +import Int. (* local env = get_environment() @@ -7,7 +7,7 @@ Import Int. assert(is_kernel_object(o1)) assert(o1) assert(o1:is_builtin()) - assert(o1:keyword() == "Builtin") + assert(o1:keyword() == "builtin") assert(o1:get_name() == name("Int", "add")) local o2 = env:find_object("xyz31213") assert(not o2) diff --git a/tests/lean/lua11.lean.expected.out b/tests/lean/lua11.lean.expected.out index 455fc70ea..8ea9e7231 100644 --- a/tests/lean/lua11.lean.expected.out +++ b/tests/lean/lua11.lean.expected.out @@ -2,5 +2,5 @@ Set: pp::unicode Imported 'Int' Int::add -BuiltinSet Nat::numeral +builtinset Nat::numeral 512 diff --git a/tests/lean/lua12.lean b/tests/lean/lua12.lean index b47b5b4b9..ddad50d27 100644 --- a/tests/lean/lua12.lean +++ b/tests/lean/lua12.lean @@ -1,5 +1,5 @@ -Import Int. -Variables x y z : Int +import Int. +variables x y z : Int (* import("util.lua") @@ -11,4 +11,4 @@ Variables x y z : Int env:add_definition("sum", def) *) -Eval sum + 3 +eval sum + 3 diff --git a/tests/lean/lua13.lean b/tests/lean/lua13.lean index 3b0c7c541..a05ec946f 100644 --- a/tests/lean/lua13.lean +++ b/tests/lean/lua13.lean @@ -1,6 +1,6 @@ -Import Int. -Variables x y z : Int -Variable f : Int -> Int -> Int +import Int. +variables x y z : Int +variable f : Int -> Int -> Int (* local t = parse_lean("fun w, f w (f y 0)") diff --git a/tests/lean/lua14.lean b/tests/lean/lua14.lean index 02f44a759..03127b606 100644 --- a/tests/lean/lua14.lean +++ b/tests/lean/lua14.lean @@ -1,6 +1,6 @@ -Import Int. -Variables x y z : Int -Variable f : Int -> Int -> Int +import Int. +variables x y z : Int +variable f : Int -> Int -> Int (* local t = parse_lean("fun w, f w (f y 0)") @@ -12,11 +12,11 @@ Variable f : Int -> Int -> Int local env = get_environment() assert(env:find_object("Int"):get_name() == name("Int")) parse_lean_cmds([[ - Eval 10 + 20 - Check x + y - Variable g : Int -> Int + eval 10 + 20 + check x + y + variable g : Int -> Int ]]) *) -Check g (f x 10) +check g (f x 10) diff --git a/tests/lean/lua15.lean b/tests/lean/lua15.lean index 940d6daa1..1a678e5e2 100644 --- a/tests/lean/lua15.lean +++ b/tests/lean/lua15.lean @@ -1,6 +1,6 @@ -Import Int. -Variables i j : Int -Variable p : Bool +import Int. +variables i j : Int +variable p : Bool (* local env = get_environment() diff --git a/tests/lean/lua16.lean b/tests/lean/lua16.lean index bca712a20..072a3b92c 100644 --- a/tests/lean/lua16.lean +++ b/tests/lean/lua16.lean @@ -1,5 +1,5 @@ -Import Int. -Variables a b : Int +import Int. +variables a b : Int (* local ios = io_state() diff --git a/tests/lean/lua17.lean b/tests/lean/lua17.lean index b90c2b220..1ff3333ec 100644 --- a/tests/lean/lua17.lean +++ b/tests/lean/lua17.lean @@ -1,6 +1,6 @@ -Import Int. -Variables a b : Int -print Options +import Int. +variables a b : Int +print options (* local ios = io_state() @@ -10,5 +10,5 @@ print Options print(parse_lean("fun x, a + x")) print(get_options()) *) -print Options -print Environment 2 \ No newline at end of file +print options +print environment 2 \ No newline at end of file diff --git a/tests/lean/lua17.lean.expected.out b/tests/lean/lua17.lean.expected.out index c1c754053..c9708ba44 100644 --- a/tests/lean/lua17.lean.expected.out +++ b/tests/lean/lua17.lean.expected.out @@ -10,5 +10,5 @@ Int::add a b λ x : ℤ, a + x ⟨pp::unicode ↦ true, pp::colors ↦ false⟩ ⟨pp::unicode ↦ true, pp::colors ↦ false⟩ -Variable a : ℤ -Variable b : ℤ +variable a : ℤ +variable b : ℤ diff --git a/tests/lean/lua18.lean b/tests/lean/lua18.lean index 05d1cb149..739fc9992 100644 --- a/tests/lean/lua18.lean +++ b/tests/lean/lua18.lean @@ -1,4 +1,4 @@ -Import Int. +import Int. (* macro("MyMacro", { macro_arg.Expr, macro_arg.Comma, macro_arg.Expr }, function (env, e1, e2) @@ -22,4 +22,4 @@ print (MyMacro 10, 20) + 20 print (Sum) print Sum 10 20 30 40 print fun x, Sum x 10 x 20 -Eval (fun x, Sum x 10 x 20) 100 +eval (fun x, Sum x 10 x 20) 100 diff --git a/tests/lean/lua2.lean b/tests/lean/lua2.lean index 7bd2c49be..05070da9a 100644 --- a/tests/lean/lua2.lean +++ b/tests/lean/lua2.lean @@ -1,4 +1,4 @@ -Variable x : Bool +variable x : Bool (* a = {} @@ -11,4 +11,4 @@ Variable x : Bool rint ("ok") *) -Variable y : Bool +variable y : Bool diff --git a/tests/lean/lua3.lean b/tests/lean/lua3.lean index 52c8a28ae..c206d843c 100644 --- a/tests/lean/lua3.lean +++ b/tests/lean/lua3.lean @@ -1,5 +1,5 @@ -Import Int. -Variable x : Int +import Int. +variable x : Int (* dofile("script.lua") diff --git a/tests/lean/lua4.lean b/tests/lean/lua4.lean index 460f64b16..bc8660b3c 100644 --- a/tests/lean/lua4.lean +++ b/tests/lean/lua4.lean @@ -1,5 +1,5 @@ -Import Int. -Variable x : Int +import Int. +variable x : Int (* -- Add a variable to the environment using Lua @@ -11,4 +11,4 @@ print("type of x is " .. tostring(typeofx)) env:add_var("y", typeofx) *) -Check x + y +check x + y diff --git a/tests/lean/lua5.lean b/tests/lean/lua5.lean index 569e9b12f..f56bab9eb 100644 --- a/tests/lean/lua5.lean +++ b/tests/lean/lua5.lean @@ -1,5 +1,5 @@ -Import Int. -Variable x : Int +import Int. +variable x : Int (* local N = 100 @@ -11,5 +11,5 @@ for i = 1, N do end *) -print Environment 101 -Check x + y_1 + y_2 \ No newline at end of file +print environment 101 +check x + y_1 + y_2 \ No newline at end of file diff --git a/tests/lean/lua5.lean.expected.out b/tests/lean/lua5.lean.expected.out index 4f351e94c..bc01956fa 100644 --- a/tests/lean/lua5.lean.expected.out +++ b/tests/lean/lua5.lean.expected.out @@ -2,105 +2,105 @@ Set: pp::unicode Imported 'Int' Assumed: x -Variable x : ℤ -Variable y_1 : ℤ -Variable y_2 : ℤ -Variable y_3 : ℤ -Variable y_4 : ℤ -Variable y_5 : ℤ -Variable y_6 : ℤ -Variable y_7 : ℤ -Variable y_8 : ℤ -Variable y_9 : ℤ -Variable y_10 : ℤ -Variable y_11 : ℤ -Variable y_12 : ℤ -Variable y_13 : ℤ -Variable y_14 : ℤ -Variable y_15 : ℤ -Variable y_16 : ℤ -Variable y_17 : ℤ -Variable y_18 : ℤ -Variable y_19 : ℤ -Variable y_20 : ℤ -Variable y_21 : ℤ -Variable y_22 : ℤ -Variable y_23 : ℤ -Variable y_24 : ℤ -Variable y_25 : ℤ -Variable y_26 : ℤ -Variable y_27 : ℤ -Variable y_28 : ℤ -Variable y_29 : ℤ -Variable y_30 : ℤ -Variable y_31 : ℤ -Variable y_32 : ℤ -Variable y_33 : ℤ -Variable y_34 : ℤ -Variable y_35 : ℤ -Variable y_36 : ℤ -Variable y_37 : ℤ -Variable y_38 : ℤ -Variable y_39 : ℤ -Variable y_40 : ℤ -Variable y_41 : ℤ -Variable y_42 : ℤ -Variable y_43 : ℤ -Variable y_44 : ℤ -Variable y_45 : ℤ -Variable y_46 : ℤ -Variable y_47 : ℤ -Variable y_48 : ℤ -Variable y_49 : ℤ -Variable y_50 : ℤ -Variable y_51 : ℤ -Variable y_52 : ℤ -Variable y_53 : ℤ -Variable y_54 : ℤ -Variable y_55 : ℤ -Variable y_56 : ℤ -Variable y_57 : ℤ -Variable y_58 : ℤ -Variable y_59 : ℤ -Variable y_60 : ℤ -Variable y_61 : ℤ -Variable y_62 : ℤ -Variable y_63 : ℤ -Variable y_64 : ℤ -Variable y_65 : ℤ -Variable y_66 : ℤ -Variable y_67 : ℤ -Variable y_68 : ℤ -Variable y_69 : ℤ -Variable y_70 : ℤ -Variable y_71 : ℤ -Variable y_72 : ℤ -Variable y_73 : ℤ -Variable y_74 : ℤ -Variable y_75 : ℤ -Variable y_76 : ℤ -Variable y_77 : ℤ -Variable y_78 : ℤ -Variable y_79 : ℤ -Variable y_80 : ℤ -Variable y_81 : ℤ -Variable y_82 : ℤ -Variable y_83 : ℤ -Variable y_84 : ℤ -Variable y_85 : ℤ -Variable y_86 : ℤ -Variable y_87 : ℤ -Variable y_88 : ℤ -Variable y_89 : ℤ -Variable y_90 : ℤ -Variable y_91 : ℤ -Variable y_92 : ℤ -Variable y_93 : ℤ -Variable y_94 : ℤ -Variable y_95 : ℤ -Variable y_96 : ℤ -Variable y_97 : ℤ -Variable y_98 : ℤ -Variable y_99 : ℤ -Variable y_100 : ℤ +variable x : ℤ +variable y_1 : ℤ +variable y_2 : ℤ +variable y_3 : ℤ +variable y_4 : ℤ +variable y_5 : ℤ +variable y_6 : ℤ +variable y_7 : ℤ +variable y_8 : ℤ +variable y_9 : ℤ +variable y_10 : ℤ +variable y_11 : ℤ +variable y_12 : ℤ +variable y_13 : ℤ +variable y_14 : ℤ +variable y_15 : ℤ +variable y_16 : ℤ +variable y_17 : ℤ +variable y_18 : ℤ +variable y_19 : ℤ +variable y_20 : ℤ +variable y_21 : ℤ +variable y_22 : ℤ +variable y_23 : ℤ +variable y_24 : ℤ +variable y_25 : ℤ +variable y_26 : ℤ +variable y_27 : ℤ +variable y_28 : ℤ +variable y_29 : ℤ +variable y_30 : ℤ +variable y_31 : ℤ +variable y_32 : ℤ +variable y_33 : ℤ +variable y_34 : ℤ +variable y_35 : ℤ +variable y_36 : ℤ +variable y_37 : ℤ +variable y_38 : ℤ +variable y_39 : ℤ +variable y_40 : ℤ +variable y_41 : ℤ +variable y_42 : ℤ +variable y_43 : ℤ +variable y_44 : ℤ +variable y_45 : ℤ +variable y_46 : ℤ +variable y_47 : ℤ +variable y_48 : ℤ +variable y_49 : ℤ +variable y_50 : ℤ +variable y_51 : ℤ +variable y_52 : ℤ +variable y_53 : ℤ +variable y_54 : ℤ +variable y_55 : ℤ +variable y_56 : ℤ +variable y_57 : ℤ +variable y_58 : ℤ +variable y_59 : ℤ +variable y_60 : ℤ +variable y_61 : ℤ +variable y_62 : ℤ +variable y_63 : ℤ +variable y_64 : ℤ +variable y_65 : ℤ +variable y_66 : ℤ +variable y_67 : ℤ +variable y_68 : ℤ +variable y_69 : ℤ +variable y_70 : ℤ +variable y_71 : ℤ +variable y_72 : ℤ +variable y_73 : ℤ +variable y_74 : ℤ +variable y_75 : ℤ +variable y_76 : ℤ +variable y_77 : ℤ +variable y_78 : ℤ +variable y_79 : ℤ +variable y_80 : ℤ +variable y_81 : ℤ +variable y_82 : ℤ +variable y_83 : ℤ +variable y_84 : ℤ +variable y_85 : ℤ +variable y_86 : ℤ +variable y_87 : ℤ +variable y_88 : ℤ +variable y_89 : ℤ +variable y_90 : ℤ +variable y_91 : ℤ +variable y_92 : ℤ +variable y_93 : ℤ +variable y_94 : ℤ +variable y_95 : ℤ +variable y_96 : ℤ +variable y_97 : ℤ +variable y_98 : ℤ +variable y_99 : ℤ +variable y_100 : ℤ x + y_1 + y_2 : ℤ diff --git a/tests/lean/lua6.lean b/tests/lean/lua6.lean index 862102086..25a9dca8b 100644 --- a/tests/lean/lua6.lean +++ b/tests/lean/lua6.lean @@ -1,20 +1,20 @@ -Import Int. -Variable x : Int -SetOption pp::notation false +import Int. +variable x : Int +setoption pp::notation false (* print(get_options()) *) -Check x + 2 +check x + 2 (* o = get_options() o = o:update(name('lean', 'pp', 'notation'), true) set_options(o) print(get_options()) *) -Check x + 2 +check x + 2 (* set_option(name('lean', 'pp', 'notation'), false) print(get_options()) *) -Variable y : Int -Check x + y +variable y : Int +check x + y diff --git a/tests/lean/lua8.lean b/tests/lean/lua8.lean index e804dee0c..7713ad357 100644 --- a/tests/lean/lua8.lean +++ b/tests/lean/lua8.lean @@ -1,5 +1,5 @@ -Import Int. -Variable x : Int +import Int. +variable x : Int (* local env = get_environment() diff --git a/tests/lean/lua8.lean.expected.out b/tests/lean/lua8.lean.expected.out index c1ce2eeb6..cd1ff6ada 100644 --- a/tests/lean/lua8.lean.expected.out +++ b/tests/lean/lua8.lean.expected.out @@ -3,5 +3,5 @@ Imported 'Int' Assumed: x x : ℤ, y : ℤ -Variable x : ℤ +variable x : ℤ nil diff --git a/tests/lean/lua9.lean b/tests/lean/lua9.lean index d383d9233..be0ae5ce5 100644 --- a/tests/lean/lua9.lean +++ b/tests/lean/lua9.lean @@ -1,5 +1,5 @@ -Import Int. -Variable x : Bool +import Int. +variable x : Bool (* import("util.lua") @@ -33,6 +33,6 @@ Variable x : Bool env:add_definition("sum1", s) *) -print Environment 1 -Eval sum1 -Variable y : Bool +print environment 1 +eval sum1 +variable y : Bool diff --git a/tests/lean/lua9.lean.expected.out b/tests/lean/lua9.lean.expected.out index bb4aa9cbd..68885517f 100644 --- a/tests/lean/lua9.lean.expected.out +++ b/tests/lean/lua9.lean.expected.out @@ -8,6 +8,6 @@ x1 + x2 ℤ → ℤ x1 + x1 + x1 + x2 + x2 ℤ -Definition sum1 : ℤ := x1 + x1 + x1 + x2 + x2 +definition sum1 : ℤ := x1 + x1 + x1 + x2 + x2 x1 + x1 + x1 + x2 + x2 Assumed: y diff --git a/tests/lean/mod1.lean b/tests/lean/mod1.lean index f493c63d6..71e40b48d 100644 --- a/tests/lean/mod1.lean +++ b/tests/lean/mod1.lean @@ -1,11 +1,11 @@ -Import cast -Import cast +import cast +import cast (* local env = environment() -- create new environment parse_lean_cmds([[ - Import cast - Import cast - Check @cast + import cast + import cast + check @cast ]], env) *) -Check @cast \ No newline at end of file +check @cast \ No newline at end of file diff --git a/tests/lean/nbug1.lean b/tests/lean/nbug1.lean index 89f3dc431..872b6d60c 100644 --- a/tests/lean/nbug1.lean +++ b/tests/lean/nbug1.lean @@ -1 +1 @@ -Notation 100 ++ : and +notation 100 ++ : and diff --git a/tests/lean/norm1.lean b/tests/lean/norm1.lean index 7a434f154..67c79cf58 100644 --- a/tests/lean/norm1.lean +++ b/tests/lean/norm1.lean @@ -1,17 +1,17 @@ -Variable N : Type -Variable P : N -> N -> N -> Bool -Variable a : N -Variable g : N -> N -Variable H : (N -> N -> N) -> N +variable N : Type +variable P : N -> N -> N -> Bool +variable a : N +variable g : N -> N +variable H : (N -> N -> N) -> N -Eval fun f : N -> N, (fun x y : N, g x) (f a) -Eval fun (a : N) (f : N -> N) (g : (N -> N) -> N -> N) (h : N -> N -> N), +eval fun f : N -> N, (fun x y : N, g x) (f a) +eval fun (a : N) (f : N -> N) (g : (N -> N) -> N -> N) (h : N -> N -> N), (fun (x : N) (y : N) (z : N), h x y) (g (fun x : N, f (f x)) (f a)) (f a) -Eval fun (a b : N) (g : Bool -> N), (fun x y : Bool, g x) (a == b) -Eval fun (a : Type) (b : a -> Type) (g : (Type U) -> Bool), (fun x y : (Type U), g x) (Pi x : a, b x) -Eval fun f : N -> N, (fun x y z : N, g x) (f a) -Eval fun f g : N -> N, (fun x y z : N, g x) (f a) -Eval fun f : N -> N, (fun x : N, fun y : N, fun z : N, P x y z) (f a) -Eval fun (f : N -> N) (a : N), (fun x : N, fun y : N, fun z : N, P x y z) (f a) -Eval fun f g : N -> N, (fun x y1 z1 : N, H ((fun x y2 z2 : N, g x) x)) (f a) -Check fun f g : N -> N, (fun x y1 z1 : N, H ((fun x y2 z2 : N, g x) x)) (f a) +eval fun (a b : N) (g : Bool -> N), (fun x y : Bool, g x) (a == b) +eval fun (a : Type) (b : a -> Type) (g : (Type U) -> Bool), (fun x y : (Type U), g x) (Pi x : a, b x) +eval fun f : N -> N, (fun x y z : N, g x) (f a) +eval fun f g : N -> N, (fun x y z : N, g x) (f a) +eval fun f : N -> N, (fun x : N, fun y : N, fun z : N, P x y z) (f a) +eval fun (f : N -> N) (a : N), (fun x : N, fun y : N, fun z : N, P x y z) (f a) +eval fun f g : N -> N, (fun x y1 z1 : N, H ((fun x y2 z2 : N, g x) x)) (f a) +check fun f g : N -> N, (fun x y1 z1 : N, H ((fun x y2 z2 : N, g x) x)) (f a) diff --git a/tests/lean/norm_bug1.lean b/tests/lean/norm_bug1.lean index 64f8a3418..9a89a4d96 100644 --- a/tests/lean/norm_bug1.lean +++ b/tests/lean/norm_bug1.lean @@ -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) (a : A) (b : A') (L2 : A' == A), @@ -9,7 +9,7 @@ Check fun (A A': TypeM) L3 : b == b' := CastEq L2 b in L3 -Check fun (A A': TypeM) +check fun (A A': TypeM) (B : A -> TypeM) (B' : A' -> TypeM) (f : Pi x : A, B x) diff --git a/tests/lean/norm_tac.lean b/tests/lean/norm_tac.lean index 3c3ef6433..5a5858230 100644 --- a/tests/lean/norm_tac.lean +++ b/tests/lean/norm_tac.lean @@ -1,13 +1,13 @@ -Import Int -Import tactic -SetOption pp::implicit true -SetOption pp::coercion true -SetOption pp::notation false -Variable vector (A : Type) (sz : Nat) : Type -Variable read {A : Type} {sz : Nat} (v : vector A sz) (i : Nat) (H : i < sz) : A -Variable V1 : vector Int 10 -Definition D := read V1 1 (by trivial) -print Environment 1 -Variable b : Bool -Definition a := b -Theorem T : b => a := (by (* imp_tac() .. normalize_tac() .. assumption_tac() *)) +import Int +import tactic +setoption pp::implicit true +setoption pp::coercion true +setoption pp::notation false +variable vector (A : Type) (sz : Nat) : Type +variable read {A : Type} {sz : Nat} (v : vector A sz) (i : Nat) (H : i < sz) : A +variable V1 : vector Int 10 +definition D := read V1 1 (by trivial) +print environment 1 +variable b : Bool +definition a := b +theorem T : b => a := (by (* imp_tac() .. normalize_tac() .. assumption_tac() *)) diff --git a/tests/lean/norm_tac.lean.expected.out b/tests/lean/norm_tac.lean.expected.out index d903defa2..67fd0fabb 100644 --- a/tests/lean/norm_tac.lean.expected.out +++ b/tests/lean/norm_tac.lean.expected.out @@ -9,7 +9,7 @@ Assumed: read Assumed: V1 Defined: D -Definition D : ℤ := @read ℤ 10 V1 1 Trivial +definition D : ℤ := @read ℤ 10 V1 1 Trivial Assumed: b Defined: a Proved: T diff --git a/tests/lean/ns1.lean b/tests/lean/ns1.lean index 50f3293ea..56512ebcb 100644 --- a/tests/lean/ns1.lean +++ b/tests/lean/ns1.lean @@ -1,19 +1,19 @@ -Import Int. +import Int. -Namespace foo. - Variable a : Nat. - Definition b := a. - Theorem T : a = b := Refl a. - Axiom H : b >= a. - Namespace bla. - Variables a c d : Int. - Check a + b + c. - EndNamespace. -EndNamespace. +namespace foo. + variable a : Nat. + definition b := a. + theorem T : a = b := Refl a. + axiom H : b >= a. + namespace bla. + variables a c d : Int. + check a + b + c. + end. +end. -Check foo::T. -Check foo::H. -Check foo::a. -Check foo::bla::a. +check foo::T. +check foo::H. +check foo::a. +check foo::bla::a. -EndNamespace. \ No newline at end of file +end \ No newline at end of file diff --git a/tests/lean/ns1.lean.expected.out b/tests/lean/ns1.lean.expected.out index 185ddf088..d3c5f7e1e 100644 --- a/tests/lean/ns1.lean.expected.out +++ b/tests/lean/ns1.lean.expected.out @@ -13,4 +13,4 @@ foo::T : foo::a = foo::b foo::H : foo::b ≥ foo::a foo::a : ℕ foo::bla::a : ℤ -Error (line: 19, pos: 0) invalid end namespace command, there are no open namespaces +Error (line: 19, pos: 0) invalid 'end', not inside of a scope or namespace diff --git a/tests/lean/overload1.lean b/tests/lean/overload1.lean index 229d59a65..d5842d480 100644 --- a/tests/lean/overload1.lean +++ b/tests/lean/overload1.lean @@ -1,12 +1,12 @@ -Variable N : Type -Variable f : Bool -> Bool -> Bool -Variable g : N -> N -> N -Infixl 10 ++ : f -Infixl 10 ++ : g +variable N : Type +variable f : Bool -> Bool -> Bool +variable g : N -> N -> N +infixl 10 ++ : f +infixl 10 ++ : g print true ++ false ++ true -SetOption lean::pp::notation false +setoption lean::pp::notation false print true ++ false ++ true -Variable a : N -Variable b : N +variable a : N +variable b : N print a ++ b ++ a print true ++ false ++ false \ No newline at end of file diff --git a/tests/lean/overload2.lean b/tests/lean/overload2.lean index da6cf5dcc..cec84b0b2 100644 --- a/tests/lean/overload2.lean +++ b/tests/lean/overload2.lean @@ -1,21 +1,21 @@ -Import Int -Import Real +import Int +import Real print 1 + true -Variable R : Type -Variable T : Type -Variable r2t : R -> T -Coercion r2t -Variable t2r : T -> R -Coercion t2r -Variable f : T -> R -> T -Variable a : T -Variable b : R -SetOption lean::pp::coercion true -SetOption lean::pp::notation false +variable R : Type +variable T : Type +variable r2t : R -> T +coercion r2t +variable t2r : T -> R +coercion t2r +variable f : T -> R -> T +variable a : T +variable b : R +setoption lean::pp::coercion true +setoption lean::pp::notation false print f a b print f b a -Variable g : R -> T -> R -Infix 10 ++ : f -Infix 10 ++ : g +variable g : R -> T -> R +infix 10 ++ : f +infix 10 ++ : g print a ++ b print b ++ a diff --git a/tests/lean/pr1.lean b/tests/lean/pr1.lean index 2f612f343..439c0650a 100644 --- a/tests/lean/pr1.lean +++ b/tests/lean/pr1.lean @@ -1,6 +1,6 @@ (* import("tactic.lua") *) -Theorem T (C A B : Bool) : C -> A -> B -> A. +theorem T (C A B : Bool) : C -> A -> B -> A. exact. done. -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/pr1.lean.expected.out b/tests/lean/pr1.lean.expected.out index 88791af06..ec238831a 100644 --- a/tests/lean/pr1.lean.expected.out +++ b/tests/lean/pr1.lean.expected.out @@ -1,4 +1,4 @@ Set: pp::colors Set: pp::unicode Proved: T -Theorem T (C A B : Bool) (H : C) (H : A) (H::1 : B) : A := H +theorem T (C A B : Bool) (H : C) (H : A) (H::1 : B) : A := H diff --git a/tests/lean/push.lean b/tests/lean/push.lean index 486241470..1387f9749 100644 --- a/tests/lean/push.lean +++ b/tests/lean/push.lean @@ -1,22 +1,22 @@ -Import Int. +import Int. -Variable first : Bool +variable first : Bool -Push - Variables a b c : Int - Variable f : Int -> Int - Eval f a -Pop +scope + variables a b c : Int + variable f : Int -> Int + eval f a +pop::scope -Eval f a -- should produce an error +eval f a -- should produce an error -print Environment 1 +print environment 1 -Push - Infixl 100 ++ : Int::add - Check 10 ++ 20 -Pop +scope + infixl 100 ++ : Int::add + check 10 ++ 20 +pop::scope -Check 10 ++ 20 -- should produce an error +check 10 ++ 20 -- should produce an error -Pop -- should produce an error \ No newline at end of file +pop::scope -- should produce an error \ No newline at end of file diff --git a/tests/lean/push.lean.expected.out b/tests/lean/push.lean.expected.out index a8252a2d5..9444e0128 100644 --- a/tests/lean/push.lean.expected.out +++ b/tests/lean/push.lean.expected.out @@ -8,7 +8,7 @@ Assumed: f f a Error (line: 11, pos: 5) unknown identifier 'f' -Variable first : Bool +variable first : Bool 10 ++ 20 : ℤ Error (line: 20, pos: 9) unknown identifier '++' Error (line: 22, pos: 0) main scope cannot be removed diff --git a/tests/lean/refute1.lean b/tests/lean/refute1.lean index aeb593c09..bb9d4a83a 100644 --- a/tests/lean/refute1.lean +++ b/tests/lean/refute1.lean @@ -1,3 +1,3 @@ -Variables a b : Bool -Axiom H : a /\ b -Theorem T : a := Refute (fun R, Absurd (Conjunct1 H) R) +variables a b : Bool +axiom H : a /\ b +theorem T : a := Refute (fun R, Absurd (Conjunct1 H) R) diff --git a/tests/lean/revapp.lean b/tests/lean/revapp.lean index 5d878c95b..44f8e4600 100644 --- a/tests/lean/revapp.lean +++ b/tests/lean/revapp.lean @@ -1,30 +1,30 @@ -Import Int. -Definition revapp {A : (Type U)} {B : A -> (Type U)} (a : A) (f : Pi (x : A), B x) : (B a) := f a. -Infixl 100 |> : revapp +import Int. +definition revapp {A : (Type U)} {B : A -> (Type U)} (a : A) (f : Pi (x : A), B x) : (B a) := f a. +infixl 100 |> : revapp -Eval 10 |> (fun x, x + 1) +eval 10 |> (fun x, x + 1) |> (fun x, x + 2) |> (fun x, 2 * x) |> (fun x, 3 - x) |> (fun x, x + 2) -Definition revcomp {A B C: (Type U)} (f : A -> B) (g : B -> C) : A -> C := +definition revcomp {A B C: (Type U)} (f : A -> B) (g : B -> C) : A -> C := fun x, g (f x) -Infixl 100 #> : revcomp +infixl 100 #> : revcomp -Eval (fun x, x + 1) #> +eval (fun x, x + 1) #> (fun x, 2 * x * x) #> (fun x, 10 + x) -Definition simple := (fun x, x + 1) #> +definition simple := (fun x, x + 1) #> (fun x, 2 * x * x) #> (fun x, 10 + x) -Check simple -Eval simple 10 +check simple +eval simple 10 -Definition simple2 := (fun x : Int, x + 1) #> +definition simple2 := (fun x : Int, x + 1) #> (fun x, 2 * x * x) #> (fun x, 10 + x) -Check simple2 -Eval simple2 (-10) +check simple2 +eval simple2 (-10) diff --git a/tests/lean/scope.lean b/tests/lean/scope.lean index 278e9800d..d3261f262 100644 --- a/tests/lean/scope.lean +++ b/tests/lean/scope.lean @@ -1,27 +1,27 @@ -Import Int. +import Int. -Scope - Variable A : Type - Variable B : Type - Variable f : A -> A -> A - Definition g (x y : A) : A := f y x - Variable h : A -> B - Variable hinv : B -> A - Axiom Inv (x : A) : hinv (h x) = x - Axiom H1 (x y : A) : f x y = f y x - Theorem f_eq_g : f = g := Abst (fun x, (Abst (fun y, +scope + variable A : Type + variable B : Type + variable f : A -> A -> A + definition g (x y : A) : A := f y x + variable h : A -> B + variable hinv : B -> A + axiom Inv (x : A) : hinv (h x) = x + axiom H1 (x y : A) : f x y = f y x + theorem f_eq_g : f = g := Abst (fun x, (Abst (fun y, let L1 : f x y = f y x := H1 x y, L2 : f y x = g x y := Refl (g x y) in Trans L1 L2))) - Theorem Inj (x y : A) (H : h x = h y) : x = y := + theorem Inj (x y : A) (H : h x = h y) : x = y := let L1 : hinv (h x) = hinv (h y) := Congr2 hinv H, L2 : hinv (h x) = x := Inv x, L3 : hinv (h y) = y := Inv y, L4 : x = hinv (h x) := Symm L2, L5 : x = hinv (h y) := Trans L4 L1 in Trans L5 L3. -EndScope +end -print Environment 3. -Eval g Int Int::sub 10 20 \ No newline at end of file +print environment 3. +eval g Int Int::sub 10 20 \ No newline at end of file diff --git a/tests/lean/scope.lean.expected.out b/tests/lean/scope.lean.expected.out index 4e71c9c2a..6d51cb115 100644 --- a/tests/lean/scope.lean.expected.out +++ b/tests/lean/scope.lean.expected.out @@ -11,12 +11,12 @@ Assumed: H1 Proved: f_eq_g Proved: Inj -Definition g (A : Type) (f : A → A → A) (x y : A) : A := f y x -Theorem f_eq_g (A : Type) (f : A → A → A) (H1 : Π x y : A, f x y = f y x) : f = g A f := +definition g (A : Type) (f : A → A → A) (x y : A) : A := f y x +theorem f_eq_g (A : Type) (f : A → A → A) (H1 : Π x y : A, f x y = f y x) : f = g A f := Abst (λ x : A, Abst (λ y : A, let L1 : f x y = f y x := H1 x y, L2 : f y x = g A f x y := Refl (g A f x y) in Trans L1 L2)) -Theorem Inj (A B : Type) (h : A → B) (hinv : B → A) (Inv : Π x : A, hinv (h x) = x) (x y : A) (H : h x = h y) : x = +theorem Inj (A B : Type) (h : A → B) (hinv : B → A) (Inv : Π x : A, hinv (h x) = x) (x y : A) (H : h x = h y) : x = y := let L1 : hinv (h x) = hinv (h y) := Congr2 hinv H, L2 : hinv (h x) = x := Inv x, diff --git a/tests/lean/showenv.l b/tests/lean/showenv.l index 893adff90..033b27078 100644 --- a/tests/lean/showenv.l +++ b/tests/lean/showenv.l @@ -1,3 +1,3 @@ -SetOption pp::colors false +setoption pp::colors false print "===BEGIN ENVIRONMENT===" -print Environment +print environment diff --git a/tests/lean/simple.lean b/tests/lean/simple.lean index 1a43aeb9f..8d9e40fe4 100644 --- a/tests/lean/simple.lean +++ b/tests/lean/simple.lean @@ -1,3 +1,3 @@ -Import Int. -Variable x : Int -Variable y : Int \ No newline at end of file +import Int. +variable x : Int +variable y : Int \ No newline at end of file diff --git a/tests/lean/single.lean b/tests/lean/single.lean index 5ffa93ba6..a9b7eca42 100644 --- a/tests/lean/single.lean +++ b/tests/lean/single.lean @@ -1,7 +1,7 @@ -Import Int. -Variables a b c : Int. +import Int. +variables a b c : Int. print a + b + c. -Check a + b. -Exit. +check a + b. +exit. -- the following line should be executed -Check a + true. +check a + true. diff --git a/tests/lean/slow/config.lean b/tests/lean/slow/config.lean index a24d59dc4..9386ad12f 100644 --- a/tests/lean/slow/config.lean +++ b/tests/lean/slow/config.lean @@ -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 diff --git a/tests/lean/slow/deep.lean b/tests/lean/slow/deep.lean index 9ffe34312..78d9b0854 100644 --- a/tests/lean/slow/deep.lean +++ b/tests/lean/slow/deep.lean @@ -1,8 +1,8 @@ -Import Int. -Definition f1 (f : Int -> Int) (x : Int) : Int := f (f (f (f x))) -Definition f2 (f : Int -> Int) (x : Int) : Int := f1 (f1 (f1 (f1 f))) x -Definition f3 (f : Int -> Int) (x : Int) : Int := f1 (f2 (f2 f)) x -Variable f : Int -> Int. -SetOption pp::width 80. -SetOption lean::pp::max_depth 2000. -Eval f3 f 0. \ No newline at end of file +import Int. +definition f1 (f : Int -> Int) (x : Int) : Int := f (f (f (f x))) +definition f2 (f : Int -> Int) (x : Int) : Int := f1 (f1 (f1 (f1 f))) x +definition f3 (f : Int -> Int) (x : Int) : Int := f1 (f2 (f2 f)) x +variable f : Int -> Int. +setoption pp::width 80. +setoption lean::pp::max_depth 2000. +eval f3 f 0. \ No newline at end of file diff --git a/tests/lean/slow/tactic1.lean b/tests/lean/slow/tactic1.lean index 1c61a87db..397612aca 100644 --- a/tests/lean/slow/tactic1.lean +++ b/tests/lean/slow/tactic1.lean @@ -1,6 +1,6 @@ -Import Int. -Definition double {A : Type} (f : A -> A) : A -> A := fun x, f (f x). -Definition big {A : Type} (f : A -> A) : A -> A := (double (double (double (double (double (double (double f))))))). +import Int. +definition double {A : Type} (f : A -> A) : A -> A := fun x, f (f x). +definition big {A : Type} (f : A -> A) : A -> A := (double (double (double (double (double (double (double f))))))). (* @@ -20,14 +20,14 @@ lazy_tac = OrElse(Then(Try(unfold_tac("eq")), congr_tac, now_tac()), *) -Theorem T1 (a b : Int) (f : Int -> Int) (H : a = b) : (big f a) = (big f b). +theorem T1 (a b : Int) (f : Int -> Int) (H : a = b) : (big f a) = (big f b). eager_tac. done. -Theorem T2 (a b : Int) (f : Int -> Int) (H : a = b) : (big f a) = (big f b). +theorem T2 (a b : Int) (f : Int -> Int) (H : a = b) : (big f a) = (big f b). lazy_tac. done. -Theorem T3 (a b : Int) (f : Int -> Int) (H : a = b) : (big f a) = ((double (double (double (double (double (double (double f))))))) b). +theorem T3 (a b : Int) (f : Int -> Int) (H : a = b) : (big f a) = ((double (double (double (double (double (double (double f))))))) b). lazy_tac. done. diff --git a/tests/lean/subst.lean b/tests/lean/subst.lean index 3259ae2d9..112bf8870 100644 --- a/tests/lean/subst.lean +++ b/tests/lean/subst.lean @@ -1,10 +1,10 @@ -Import Int. -Variable a : Int -Variable n : Nat -Axiom H1 : a + a + a = 10 -Axiom H2 : a = n -Theorem T : a + n + a = 10 := Subst H1 H2 -SetOption pp::coercion true -SetOption pp::notation false -SetOption pp::implicit true -print Environment 1. +import Int. +variable a : Int +variable n : Nat +axiom H1 : a + a + a = 10 +axiom H2 : a = n +theorem T : a + n + a = 10 := Subst H1 H2 +setoption pp::coercion true +setoption pp::notation false +setoption pp::implicit true +print environment 1. diff --git a/tests/lean/subst.lean.expected.out b/tests/lean/subst.lean.expected.out index 846790642..c69a7aa25 100644 --- a/tests/lean/subst.lean.expected.out +++ b/tests/lean/subst.lean.expected.out @@ -9,5 +9,5 @@ Set: lean::pp::coercion Set: lean::pp::notation Set: lean::pp::implicit -Theorem T : @eq ℤ (Int::add (Int::add a (nat_to_int n)) a) (nat_to_int 10) := +theorem T : @eq ℤ (Int::add (Int::add a (nat_to_int n)) a) (nat_to_int 10) := @Subst ℤ a (nat_to_int n) (λ x : ℤ, @eq ℤ (Int::add (Int::add a x) a) (nat_to_int 10)) H1 H2 diff --git a/tests/lean/subst2.lean b/tests/lean/subst2.lean index 75c155e0f..5d70ee285 100644 --- a/tests/lean/subst2.lean +++ b/tests/lean/subst2.lean @@ -1,5 +1,5 @@ (* import("tactic.lua") *) -Theorem T (A : Type) (p : A -> Bool) (f : A -> A -> A) : forall x y z, p (f x x) => x = y => x = z => p (f y z). +theorem T (A : Type) (p : A -> Bool) (f : A -> A -> A) : forall x y z, p (f x x) => x = y => x = z => p (f y z). apply ForallIntro. beta. apply ForallIntro. @@ -12,4 +12,4 @@ Theorem T (A : Type) (p : A -> Bool) (f : A -> A -> A) : forall x y z, p (f x x) apply (Subst (Subst H H::1) H::2) done. -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/subst2.lean.expected.out b/tests/lean/subst2.lean.expected.out index d535e77de..325242d3b 100644 --- a/tests/lean/subst2.lean.expected.out +++ b/tests/lean/subst2.lean.expected.out @@ -1,7 +1,7 @@ Set: pp::colors Set: pp::unicode Proved: T -Theorem T (A : Type) (p : A → Bool) (f : A → A → A) : ∀ x y z : A, +theorem T (A : Type) (p : A → Bool) (f : A → A → A) : ∀ x y z : A, p (f x x) ⇒ x = y ⇒ x = z ⇒ p (f y z) := ForallIntro (λ x : A, diff --git a/tests/lean/subst3.lean b/tests/lean/subst3.lean index 850946cf4..b9ae590fd 100644 --- a/tests/lean/subst3.lean +++ b/tests/lean/subst3.lean @@ -1,7 +1,7 @@ (* import("macros.lua") *) -Theorem T (A : Type) (p : A -> Bool) (f : A -> A -> A) : forall x y z, p (f x x) => x = y => x = z => p (f y z) := +theorem T (A : Type) (p : A -> Bool) (f : A -> A -> A) : forall x y z, p (f x x) => x = y => x = z => p (f y z) := take x y z, Assume (H1 : p (f x x)) (H2 : x = y) (H3 : x = z), Subst' H1 H2 H3. -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/subst3.lean.expected.out b/tests/lean/subst3.lean.expected.out index 3714c1180..5ad8b09d0 100644 --- a/tests/lean/subst3.lean.expected.out +++ b/tests/lean/subst3.lean.expected.out @@ -1,7 +1,7 @@ Set: pp::colors Set: pp::unicode Proved: T -Theorem T (A : Type) (p : A → Bool) (f : A → A → A) : ∀ x y z : A, +theorem T (A : Type) (p : A → Bool) (f : A → A → A) : ∀ x y z : A, p (f x x) ⇒ x = y ⇒ x = z ⇒ p (f y z) := ForallIntro (λ x : A, diff --git a/tests/lean/tactic1.lean b/tests/lean/tactic1.lean index d4212d2db..278f420a3 100644 --- a/tests/lean/tactic1.lean +++ b/tests/lean/tactic1.lean @@ -1,4 +1,4 @@ -Variables p q r : Bool +variables p q r : Bool (* local env = get_environment() @@ -8,4 +8,4 @@ Variables p q r : Bool env:add_theorem("T1", conjecture, proof) *) -print Environment 1. +print environment 1. diff --git a/tests/lean/tactic1.lean.expected.out b/tests/lean/tactic1.lean.expected.out index 3e3baf507..3710f2536 100644 --- a/tests/lean/tactic1.lean.expected.out +++ b/tests/lean/tactic1.lean.expected.out @@ -3,4 +3,4 @@ Assumed: p Assumed: q Assumed: r -Theorem T1 : p ⇒ q ⇒ p ∧ q := Discharge (λ H : p, Discharge (λ H::1 : q, Conj H H::1)) +theorem T1 : p ⇒ q ⇒ p ∧ q := Discharge (λ H : p, Discharge (λ H::1 : q, Conj H H::1)) diff --git a/tests/lean/tactic10.lean b/tests/lean/tactic10.lean index a944d5d07..f9a722c93 100644 --- a/tests/lean/tactic10.lean +++ b/tests/lean/tactic10.lean @@ -1,8 +1,8 @@ (* import("tactic.lua") *) -Definition f(a : Bool) : Bool := not a. -Definition g(a b : Bool) : Bool := a \/ b. +definition f(a : Bool) : Bool := not a. +definition g(a b : Bool) : Bool := a \/ b. -Theorem T1 (a b : Bool) : (g a b) => (f b) => a := _. +theorem T1 (a b : Bool) : (g a b) => (f b) => a := _. unfold_all apply Discharge apply Discharge @@ -15,10 +15,10 @@ Theorem T1 (a b : Bool) : (g a b) => (f b) => a := _. simple_tac = Repeat(unfold_tac()) .. Repeat(OrElse(imp_tac(), conj_hyp_tac(), assumption_tac(), absurd_tac(), conj_hyp_tac(), disj_hyp_tac())) *) -Definition h(a b : Bool) : Bool := g a b. +definition h(a b : Bool) : Bool := g a b. -Theorem T2 (a b : Bool) : (h a b) => (f b) => a := _. +theorem T2 (a b : Bool) : (h a b) => (f b) => a := _. simple_tac done. -print Environment 1. +print environment 1. diff --git a/tests/lean/tactic10.lean.expected.out b/tests/lean/tactic10.lean.expected.out index 1cdd18855..48eb3b860 100644 --- a/tests/lean/tactic10.lean.expected.out +++ b/tests/lean/tactic10.lean.expected.out @@ -5,5 +5,5 @@ Proved: T1 Defined: h Proved: T2 -Theorem T2 (a b : Bool) : h a b ⇒ f b ⇒ a := +theorem T2 (a b : Bool) : h a b ⇒ f b ⇒ a := Discharge (λ H : a ∨ b, Discharge (λ H::1 : ¬ b, DisjCases H (λ H : a, H) (λ H : b, AbsurdElim a H H::1))) diff --git a/tests/lean/tactic11.lean b/tests/lean/tactic11.lean index 87a3996f5..742f9640b 100644 --- a/tests/lean/tactic11.lean +++ b/tests/lean/tactic11.lean @@ -1,5 +1,5 @@ (* import("tactic.lua") *) -Theorem T (a b : Bool) : ((fun x, x /\ b) a) => ((fun x, x) a) := _ . +theorem T (a b : Bool) : ((fun x, x /\ b) a) => ((fun x, x) a) := _ . beta. apply Discharge. conj_hyp. diff --git a/tests/lean/tactic12.lean b/tests/lean/tactic12.lean index 172b67fd4..b3218dfbe 100644 --- a/tests/lean/tactic12.lean +++ b/tests/lean/tactic12.lean @@ -1,13 +1,13 @@ (* import("tactic.lua") *) -Theorem T (a b : Bool) : ((fun x, x /\ b) a) => ((fun x, x) a). +theorem T (a b : Bool) : ((fun x, x /\ b) a) => ((fun x, x) a). beta. apply Discharge. conj_hyp. exact. done. -Variables p q : Bool. -Theorem T2 : p /\ q => q. +variables p q : Bool. +theorem T2 : p /\ q => q. apply Discharge. conj_hyp. exact. diff --git a/tests/lean/tactic13.lean b/tests/lean/tactic13.lean index 39ebb8f27..5cab16f8b 100644 --- a/tests/lean/tactic13.lean +++ b/tests/lean/tactic13.lean @@ -1,6 +1,6 @@ -Import Int. +import Int. (* import("tactic.lua") *) -Variable f : Int -> Int -> Int +variable f : Int -> Int -> Int (* refl_tac = apply_tac("Refl") @@ -11,12 +11,12 @@ unfold_homo_eq_tac = unfold_tac("eq") auto = unfold_homo_eq_tac .. Repeat(OrElse(refl_tac, congr_tac, assumption_tac(), Then(symm_tac, assumption_tac(), now_tac()))) *) -Theorem T1 (a b c : Int) (H1 : a = b) (H2 : a = c) : (f (f a a) b) = (f (f b c) a). +theorem T1 (a b c : Int) (H1 : a = b) (H2 : a = c) : (f (f a a) b) = (f (f b c) a). auto. done. -Theorem T2 (a b c : Int) (H1 : a = b) (H2 : a = c) : (f (f a c)) = (f (f b a)). +theorem T2 (a b c : Int) (H1 : a = b) (H2 : a = c) : (f (f a c)) = (f (f b a)). auto. done. -print Environment 2. \ No newline at end of file +print environment 2. \ No newline at end of file diff --git a/tests/lean/tactic13.lean.expected.out b/tests/lean/tactic13.lean.expected.out index 65e28cd7c..fea2a706b 100644 --- a/tests/lean/tactic13.lean.expected.out +++ b/tests/lean/tactic13.lean.expected.out @@ -4,7 +4,7 @@ Assumed: f Proved: T1 Proved: T2 -Theorem T1 (a b c : ℤ) (H1 : a = b) (H2 : a = c) : f (f a a) b = f (f b c) a := +theorem T1 (a b c : ℤ) (H1 : a = b) (H2 : a = c) : f (f a a) b = f (f b c) a := Congr (Congr (Refl f) (Congr (Congr (Refl f) H1) H2)) (Symm H1) -Theorem T2 (a b c : ℤ) (H1 : a = b) (H2 : a = c) : f (f a c) = f (f b a) := +theorem T2 (a b c : ℤ) (H1 : a = b) (H2 : a = c) : f (f a c) = f (f b a) := Congr (Refl f) (Congr (Congr (Refl f) H1) (Symm H2)) diff --git a/tests/lean/tactic14.lean b/tests/lean/tactic14.lean index 132592700..9fced163c 100644 --- a/tests/lean/tactic14.lean +++ b/tests/lean/tactic14.lean @@ -1,4 +1,4 @@ -Import Int. +import Int. (* -- Tactic for trying to prove goal using Reflexivity, Congruence and available assumptions @@ -6,8 +6,8 @@ congr_tac = Try(unfold_tac("eq")) .. Repeat(OrElse(apply_tac("Refl"), apply_tac( *) -Theorem T1 (a b : Int) (f : Int -> Int) : a = b -> (f (f a)) = (f (f b)) := +theorem T1 (a b : Int) (f : Int -> Int) : a = b -> (f (f a)) = (f (f b)) := fun assumption : a = b, have (f (f a)) = (f (f b)) by congr_tac -print Environment 1. +print environment 1. diff --git a/tests/lean/tactic14.lean.expected.out b/tests/lean/tactic14.lean.expected.out index 41354de88..d21bf7293 100644 --- a/tests/lean/tactic14.lean.expected.out +++ b/tests/lean/tactic14.lean.expected.out @@ -2,5 +2,5 @@ Set: pp::unicode Imported 'Int' Proved: T1 -Theorem T1 (a b : ℤ) (f : ℤ → ℤ) (assumption : a = b) : f (f a) = f (f b) := +theorem T1 (a b : ℤ) (f : ℤ → ℤ) (assumption : a = b) : f (f a) = f (f b) := Congr (Refl f) (Congr (Refl f) assumption) diff --git a/tests/lean/tactic2.lean b/tests/lean/tactic2.lean index 4d255032b..a4ed118b2 100644 --- a/tests/lean/tactic2.lean +++ b/tests/lean/tactic2.lean @@ -1,7 +1,7 @@ (* import("tactic.lua") *) -Variables p q r : Bool +variables p q r : Bool -Theorem T1 : p => q => p /\ q := +theorem T1 : p => q => p /\ q := Discharge (fun H1, Discharge (fun H2, let H1 : p := _, @@ -17,22 +17,22 @@ Theorem T1 : p => q => p /\ q := simple_tac = Repeat(imp_tac() ^ conj_tac() ^ assumption_tac()) *) -Theorem T2 : p => q => p /\ q /\ p := _. +theorem T2 : p => q => p /\ q /\ p := _. simple_tac done -print Environment 1 +print environment 1 -Theorem T3 : p => p /\ q => r => q /\ r /\ p := _. +theorem T3 : p => p /\ q => r => q /\ r /\ p := _. (* Repeat(OrElse(imp_tac(), conj_tac(), conj_hyp_tac(), assumption_tac())) *) done -- Display proof term generated by previous tac -print Environment 1 +print environment 1 -Theorem T4 : p => p /\ q => r => q /\ r /\ p := _. +theorem T4 : p => p /\ q => r => q /\ r /\ p := _. Repeat (OrElse (apply Discharge) (apply Conj) conj_hyp exact) done -- Display proof term generated by previous tac -- -print Environment 1 \ No newline at end of file +print environment 1 \ No newline at end of file diff --git a/tests/lean/tactic2.lean.expected.out b/tests/lean/tactic2.lean.expected.out index e1b42a894..e71265514 100644 --- a/tests/lean/tactic2.lean.expected.out +++ b/tests/lean/tactic2.lean.expected.out @@ -5,14 +5,14 @@ Assumed: r Proved: T1 Proved: T2 -Theorem T2 : p ⇒ q ⇒ p ∧ q ∧ p := Discharge (λ H : p, Discharge (λ H::1 : q, Conj H (Conj H::1 H))) +theorem T2 : p ⇒ q ⇒ p ∧ q ∧ p := Discharge (λ H : p, Discharge (λ H::1 : q, Conj H (Conj H::1 H))) Proved: T3 -Theorem T3 : p ⇒ p ∧ q ⇒ r ⇒ q ∧ r ∧ p := +theorem T3 : p ⇒ p ∧ q ⇒ r ⇒ q ∧ r ∧ p := Discharge (λ H : p, Discharge (λ H::1 : p ∧ q, Discharge (λ H::2 : r, Conj (Conjunct2 H::1) (Conj H::2 (Conjunct1 H::1))))) Proved: T4 -Theorem T4 : p ⇒ p ∧ q ⇒ r ⇒ q ∧ r ∧ p := +theorem T4 : p ⇒ p ∧ q ⇒ r ⇒ q ∧ r ∧ p := Discharge (λ H : p, Discharge diff --git a/tests/lean/tactic3.lean b/tests/lean/tactic3.lean index 20048fa35..9e58081cc 100644 --- a/tests/lean/tactic3.lean +++ b/tests/lean/tactic3.lean @@ -1,9 +1,9 @@ (* import("tactic.lua") *) -Variables p q r : Bool +variables p q r : Bool -Theorem T1 : p => p /\ q => r => q /\ r /\ p := _. +theorem T1 : p => p /\ q => r => q /\ r /\ p := _. (* Repeat(OrElse(imp_tac(), conj_tac(), conj_hyp_tac(), assumption_tac())) *) done -- Display proof term generated by previous tactic -print Environment 1 \ No newline at end of file +print environment 1 \ No newline at end of file diff --git a/tests/lean/tactic3.lean.expected.out b/tests/lean/tactic3.lean.expected.out index 8536b12cd..362e992e9 100644 --- a/tests/lean/tactic3.lean.expected.out +++ b/tests/lean/tactic3.lean.expected.out @@ -4,7 +4,7 @@ Assumed: q Assumed: r Proved: T1 -Theorem T1 : p ⇒ p ∧ q ⇒ r ⇒ q ∧ r ∧ p := +theorem T1 : p ⇒ p ∧ q ⇒ r ⇒ q ∧ r ∧ p := Discharge (λ H : p, Discharge (λ H::1 : p ∧ q, Discharge (λ H::2 : r, Conj (Conjunct2 H::1) (Conj H::2 (Conjunct1 H::1))))) diff --git a/tests/lean/tactic4.lean b/tests/lean/tactic4.lean index 509682344..eeae9a2ba 100644 --- a/tests/lean/tactic4.lean +++ b/tests/lean/tactic4.lean @@ -2,8 +2,8 @@ simple_tac = Repeat(OrElse(imp_tac(), conj_tac())) .. assumption_tac() *) -Theorem T4 (a b : Bool) : a => b => a /\ b := _. +theorem T4 (a b : Bool) : a => b => a /\ b := _. simple_tac done -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/tactic4.lean.expected.out b/tests/lean/tactic4.lean.expected.out index 49bf244cc..3da2c3c0c 100644 --- a/tests/lean/tactic4.lean.expected.out +++ b/tests/lean/tactic4.lean.expected.out @@ -1,4 +1,4 @@ Set: pp::colors Set: pp::unicode Proved: T4 -Theorem T4 (a b : Bool) : a ⇒ b ⇒ a ∧ b := Discharge (λ H : a, Discharge (λ H::1 : b, Conj H H::1)) +theorem T4 (a b : Bool) : a ⇒ b ⇒ a ∧ b := Discharge (λ H : a, Discharge (λ H::1 : b, Conj H H::1)) diff --git a/tests/lean/tactic5.lean b/tests/lean/tactic5.lean index c0205251e..a6b9dfe16 100644 --- a/tests/lean/tactic5.lean +++ b/tests/lean/tactic5.lean @@ -2,8 +2,8 @@ simple_tac = Repeat(OrElse(imp_tac(), conj_tac())) .. assumption_tac() *) -Theorem T4 (a b : Bool) : a => b => a /\ b /\ a := _. +theorem T4 (a b : Bool) : a => b => a /\ b /\ a := _. simple_tac done -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/tactic5.lean.expected.out b/tests/lean/tactic5.lean.expected.out index da29a8912..88cdafaf7 100644 --- a/tests/lean/tactic5.lean.expected.out +++ b/tests/lean/tactic5.lean.expected.out @@ -1,5 +1,5 @@ Set: pp::colors Set: pp::unicode Proved: T4 -Theorem T4 (a b : Bool) : a ⇒ b ⇒ a ∧ b ∧ a := +theorem T4 (a b : Bool) : a ⇒ b ⇒ a ∧ b ∧ a := Discharge (λ H : a, Discharge (λ H::1 : b, Conj H (Conj H::1 H))) diff --git a/tests/lean/tactic6.lean b/tests/lean/tactic6.lean index 1a2ce641f..83723cdb4 100644 --- a/tests/lean/tactic6.lean +++ b/tests/lean/tactic6.lean @@ -1,5 +1,5 @@ (* import("tactic.lua") *) -Theorem T (a b c : Bool): a => b /\ c => c /\ a /\ b := _. +theorem T (a b c : Bool): a => b /\ c => c /\ a /\ b := _. apply Discharge apply Discharge conj_hyp @@ -8,7 +8,7 @@ Theorem T (a b c : Bool): a => b /\ c => c /\ a /\ b := _. exact done -Theorem T2 (a b c : Bool): a => b /\ c => c /\ a /\ b := _. +theorem T2 (a b c : Bool): a => b /\ c => c /\ a /\ b := _. apply Discharge apply Discharge conj_hyp diff --git a/tests/lean/tactic8.lean b/tests/lean/tactic8.lean index 5a4007f58..0bfc64bca 100644 --- a/tests/lean/tactic8.lean +++ b/tests/lean/tactic8.lean @@ -1,9 +1,9 @@ (* import("tactic.lua") *) -Theorem T (a b : Bool) : a \/ b => (not b) => a := _. +theorem T (a b : Bool) : a \/ b => (not b) => a := _. apply Discharge apply Discharge disj_hyp exact absurd done -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/tactic8.lean.expected.out b/tests/lean/tactic8.lean.expected.out index e822e23fc..98cc50f82 100644 --- a/tests/lean/tactic8.lean.expected.out +++ b/tests/lean/tactic8.lean.expected.out @@ -1,5 +1,5 @@ Set: pp::colors Set: pp::unicode Proved: T -Theorem T (a b : Bool) : a ∨ b ⇒ ¬ b ⇒ a := +theorem T (a b : Bool) : a ∨ b ⇒ ¬ b ⇒ a := Discharge (λ H : a ∨ b, Discharge (λ H::1 : ¬ b, DisjCases H (λ H : a, H) (λ H : b, AbsurdElim a H H::1))) diff --git a/tests/lean/tactic9.lean b/tests/lean/tactic9.lean index ef6241130..c151f757e 100644 --- a/tests/lean/tactic9.lean +++ b/tests/lean/tactic9.lean @@ -1,7 +1,7 @@ (* import("tactic.lua") *) -Definition f(a : Bool) : Bool := not a. +definition f(a : Bool) : Bool := not a. -Theorem T (a b : Bool) : a \/ b => (f b) => a := _. +theorem T (a b : Bool) : a \/ b => (f b) => a := _. apply Discharge apply Discharge disj_hyp @@ -10,4 +10,4 @@ Theorem T (a b : Bool) : a \/ b => (f b) => a := _. absurd done -print Environment 1. \ No newline at end of file +print environment 1. \ No newline at end of file diff --git a/tests/lean/tactic9.lean.expected.out b/tests/lean/tactic9.lean.expected.out index c20076798..d7e9fc32b 100644 --- a/tests/lean/tactic9.lean.expected.out +++ b/tests/lean/tactic9.lean.expected.out @@ -2,5 +2,5 @@ Set: pp::unicode Defined: f Proved: T -Theorem T (a b : Bool) : a ∨ b ⇒ f b ⇒ a := +theorem T (a b : Bool) : a ∨ b ⇒ f b ⇒ a := Discharge (λ H : a ∨ b, Discharge (λ H::1 : f b, DisjCases H (λ H : a, H) (λ H : b, AbsurdElim a H H::1))) diff --git a/tests/lean/tst1.lean b/tests/lean/tst1.lean index ce374b70e..c0137a4e2 100644 --- a/tests/lean/tst1.lean +++ b/tests/lean/tst1.lean @@ -1,26 +1,26 @@ -- Define a "fake" type to simulate the natural numbers. This is just a test. -Variable N : Type -Variable lt : N -> N -> Bool -Variable zero : N -Variable one : N -Variable two : N -Variable three : N -Infix 50 < : lt -Axiom two_lt_three : two < three -Definition vector (A : Type) (n : N) : Type := Pi (i : N) (H : i < n), A -Definition const {A : Type} (n : N) (d : A) : vector A n := fun (i : N) (H : i < n), d -Definition update {A : Type} {n : N} (v : vector A n) (i : N) (d : A) : vector A n := fun (j : N) (H : j < n), if (j = i) d (v j H) -Definition select {A : Type} {n : N} (v : vector A n) (i : N) (H : i < n) : A := v i H -Definition map {A B C : Type} {n : N} (f : A -> B -> C) (v1 : vector A n) (v2 : vector B n) : vector C n := fun (i : N) (H : i < n), f (v1 i H) (v2 i H) -print Environment 10 -Check select (update (const three false) two true) two two_lt_three -Eval select (update (const three false) two true) two two_lt_three -Check update (const three false) two true +variable N : Type +variable lt : N -> N -> Bool +variable zero : N +variable one : N +variable two : N +variable three : N +infix 50 < : lt +axiom two_lt_three : two < three +definition vector (A : Type) (n : N) : Type := Pi (i : N) (H : i < n), A +definition const {A : Type} (n : N) (d : A) : vector A n := fun (i : N) (H : i < n), d +definition update {A : Type} {n : N} (v : vector A n) (i : N) (d : A) : vector A n := fun (j : N) (H : j < n), if (j = i) d (v j H) +definition select {A : Type} {n : N} (v : vector A n) (i : N) (H : i < n) : A := v i H +definition map {A B C : Type} {n : N} (f : A -> B -> C) (v1 : vector A n) (v2 : vector B n) : vector C n := fun (i : N) (H : i < n), f (v1 i H) (v2 i H) +print environment 10 +check select (update (const three false) two true) two two_lt_three +eval select (update (const three false) two true) two two_lt_three +check update (const three false) two true print "\n--------" -Check @select +check @select print "\nmap type ---> " -Check @map +check @map print "\nmap normal form --> " -Eval @map +eval @map print "\nupdate normal form --> " -Eval @update +eval @update diff --git a/tests/lean/tst1.lean.expected.out b/tests/lean/tst1.lean.expected.out index 6e94bf372..eb3bc6ba3 100644 --- a/tests/lean/tst1.lean.expected.out +++ b/tests/lean/tst1.lean.expected.out @@ -12,17 +12,17 @@ Defined: update Defined: select Defined: map -Variable one : N -Variable two : N -Variable three : N -Infix 50 < : lt -Axiom two_lt_three : two < three -Definition vector (A : Type) (n : N) : Type := Π (i : N), i < n → A -Definition const {A : Type} (n : N) (d : A) : vector A n := λ (i : N) (H : i < n), d -Definition update {A : Type} {n : N} (v : vector A n) (i : N) (d : A) : vector A n := +variable one : N +variable two : N +variable three : N +infix 50 < : lt +axiom two_lt_three : two < three +definition vector (A : Type) (n : N) : Type := Π (i : N), i < n → A +definition const {A : Type} (n : N) (d : A) : vector A n := λ (i : N) (H : i < n), d +definition update {A : Type} {n : N} (v : vector A n) (i : N) (d : A) : vector A n := λ (j : N) (H : j < n), if (j = i) d (v j H) -Definition select {A : Type} {n : N} (v : vector A n) (i : N) (H : i < n) : A := v i H -Definition map {A B C : Type} {n : N} (f : A → B → C) (v1 : vector A n) (v2 : vector B n) : vector C n := +definition select {A : Type} {n : N} (v : vector A n) (i : N) (H : i < n) : A := v i H +definition map {A B C : Type} {n : N} (f : A → B → C) (v1 : vector A n) (v2 : vector B n) : vector C n := λ (i : N) (H : i < n), f (v1 i H) (v2 i H) select (update (const three ⊥) two ⊤) two two_lt_three : Bool if (two == two) ⊤ ⊥ diff --git a/tests/lean/tst10.lean b/tests/lean/tst10.lean index cb2cf313e..88f5af685 100644 --- a/tests/lean/tst10.lean +++ b/tests/lean/tst10.lean @@ -1,5 +1,5 @@ -Variable a : Bool -Variable b : Bool +variable a : Bool +variable b : Bool -- Conjunctions print a && b print a && b && a @@ -15,12 +15,12 @@ print (or a b) print or a (or a b) -- Simple Formulas print a => b => a -Check a => b -Eval a => a -Eval true => a +check a => b +eval a => a +eval true => a -- Simple proof -Axiom H1 : a -Axiom H2 : a => b -Check @MP +axiom H1 : a +axiom H2 : a => b +check @MP print MP H2 H1 -Check MP H2 H1 +check MP H2 H1 diff --git a/tests/lean/tst11.lean b/tests/lean/tst11.lean index 262d110dc..99b88af75 100644 --- a/tests/lean/tst11.lean +++ b/tests/lean/tst11.lean @@ -1,12 +1,12 @@ -Definition xor (x y : Bool) : Bool := (not x) = y -Infixr 50 ⊕ : xor +definition xor (x y : Bool) : Bool := (not x) = y +infixr 50 ⊕ : xor print xor true false -Eval xor true true -Eval xor true false -Variable a : Bool +eval xor true true +eval xor true false +variable a : Bool print a ⊕ a ⊕ a -Check @Subst -Theorem EM2 (a : Bool) : a \/ (not a) := +check @Subst +theorem EM2 (a : Bool) : a \/ (not a) := Case (fun x : Bool, x \/ (not x)) Trivial Trivial a -Check EM2 -Check EM2 a +check EM2 +check EM2 a diff --git a/tests/lean/tst12.lean b/tests/lean/tst12.lean index 6e0a6af89..ac798c02f 100644 --- a/tests/lean/tst12.lean +++ b/tests/lean/tst12.lean @@ -6,7 +6,7 @@ print let x := true, arg2 /\ arg1 <=> arg1 \/ arg2 \/ arg2) in (f x y) \/ z) -Eval let x := true, +eval let x := true, y := true, z := x /\ y, f := (fun arg1 arg2 : Bool, arg1 /\ arg2 <=> diff --git a/tests/lean/tst14.lean b/tests/lean/tst14.lean index f4a2bda39..6f865e449 100644 --- a/tests/lean/tst14.lean +++ b/tests/lean/tst14.lean @@ -1,6 +1,6 @@ -Import Int. +import Int. print Int -> Int -> Int -Variable f : Int -> Int -> Int -Eval f 0 -Check f 0 -Check f 0 1 +variable f : Int -> Int -> Int +eval f 0 +check f 0 +check f 0 1 diff --git a/tests/lean/tst15.lean b/tests/lean/tst15.lean index 6940b3ebe..5e0956c62 100644 --- a/tests/lean/tst15.lean +++ b/tests/lean/tst15.lean @@ -1,20 +1,20 @@ -Variable x : (Type max U+1+2 M+1 M+2 3) -Check x -Variable f : (Type U+10) -> Type -Check f -Check f x -Check (Type 4) -Check x -Check (Type max U M) +variable x : (Type max U+1+2 M+1 M+2 3) +check x +variable f : (Type U+10) -> Type +check f +check f x +check (Type 4) +check x +check (Type max U M) print (Type U+3) -Check (Type U+3) -Check (Type U ⊔ M) -Check (Type U ⊔ M ⊔ 3) +check (Type U+3) +check (Type U ⊔ M) +check (Type U ⊔ M ⊔ 3) print (Type U+1 ⊔ M ⊔ 3) -Check (Type U+1 ⊔ M ⊔ 3) +check (Type U+1 ⊔ M ⊔ 3) print (Type U) -> (Type 5) -Check (Type U) -> (Type 5) -Check (Type M ⊔ 3) -> (Type U+5) +check (Type U) -> (Type 5) +check (Type M ⊔ 3) -> (Type U+5) print (Type M ⊔ 3) -> (Type U) -> (Type 5) -Check (Type M ⊔ 3) -> (Type U) -> (Type 5) +check (Type M ⊔ 3) -> (Type U) -> (Type 5) print (Type U) diff --git a/tests/lean/tst16.lean b/tests/lean/tst16.lean index 76c48160e..f42f460be 100644 --- a/tests/lean/tst16.lean +++ b/tests/lean/tst16.lean @@ -1,15 +1,15 @@ -Variable f : Type -> Bool +variable f : Type -> Bool print forall a b : Type, (f a) = (f b) -Variable g : Bool -> Bool -> Bool +variable g : Bool -> Bool -> Bool print forall (a b : Type) (c : Bool), g c ((f a) = (f b)) print exists (a b : Type) (c : Bool), g c ((f a) = (f b)) print forall (a b : Type) (c : Bool), (g c (f a) = (f b)) ⇒ (f a) -Check forall (a b : Type) (c : Bool), g c ((f a) = (f b)) +check forall (a b : Type) (c : Bool), g c ((f a) = (f b)) print ∀ (a b : Type) (c : Bool), g c ((f a) = (f b)) print ∀ a b : Type, (f a) = (f b) print ∃ a b : Type, (f a) = (f b) ∧ (f a) print ∃ a b : Type, (f a) = (f b) ∨ (f b) -Variable a : Bool +variable a : Bool print (f a) ∨ (f a) print (f a) = a ∨ (f a) print (f a) = a ∧ (f a) diff --git a/tests/lean/tst17.lean b/tests/lean/tst17.lean index 509c86d3a..6c9a2a6af 100644 --- a/tests/lean/tst17.lean +++ b/tests/lean/tst17.lean @@ -1,5 +1,5 @@ -Variable f : Type -> Bool -Variable g : Type -> Type -> Bool +variable f : Type -> Bool +variable g : Type -> Type -> Bool print forall (a b : Type), exists (c : Type), (g a b) = (f c) -Check forall (a b : Type), exists (c : Type), (g a b) = (f c) -Eval forall (a b : Type), exists (c : Type), (g a b) = (f c) +check forall (a b : Type), exists (c : Type), (g a b) = (f c) +eval forall (a b : Type), exists (c : Type), (g a b) = (f c) diff --git a/tests/lean/tst2.lean b/tests/lean/tst2.lean index 7aa3ba63b..49e28724f 100644 --- a/tests/lean/tst2.lean +++ b/tests/lean/tst2.lean @@ -1,11 +1,11 @@ -print Options -Variable a : Bool -Variable b : Bool +print options +variable a : Bool +variable b : Bool print a/\b -SetOption lean::pp::notation false -print Options +setoption lean::pp::notation false +print options print a/\b -print Environment 2 -SetOption lean::pp::notation true -print Options +print environment 2 +setoption lean::pp::notation true +print options print a/\b diff --git a/tests/lean/tst2.lean.expected.out b/tests/lean/tst2.lean.expected.out index e685a1b71..0c0e587f7 100644 --- a/tests/lean/tst2.lean.expected.out +++ b/tests/lean/tst2.lean.expected.out @@ -7,8 +7,8 @@ a ∧ b Set: lean::pp::notation ⟨lean::pp::notation ↦ false, pp::unicode ↦ true, pp::colors ↦ false⟩ and a b -Variable a : Bool -Variable b : Bool +variable a : Bool +variable b : Bool Set: lean::pp::notation ⟨lean::pp::notation ↦ true, pp::unicode ↦ true, pp::colors ↦ false⟩ a ∧ b diff --git a/tests/lean/tst3.lean b/tests/lean/tst3.lean index 310f2d0b5..dadf5f1fa 100644 --- a/tests/lean/tst3.lean +++ b/tests/lean/tst3.lean @@ -1,58 +1,58 @@ -SetOption lean::parser::verbose false. -Notation 10 if _ then _ : implies. -print Environment 1. +setoption lean::parser::verbose false. +notation 10 if _ then _ : implies. +print environment 1. print if true then false. -Variable a : Bool. +variable a : Bool. print if true then if a then false. -SetOption lean::pp::notation false. +setoption lean::pp::notation false. print if true then if a then false. -Variable A : Type. -Variable f : A -> A -> A -> Bool. -Notation 100 _ |- _ ; _ : f. -print Environment 1. -Variable c : A. -Variable d : A. -Variable e : A. +variable A : Type. +variable f : A -> A -> A -> Bool. +notation 100 _ |- _ ; _ : f. +print environment 1. +variable c : A. +variable d : A. +variable e : A. print c |- d ; e. -SetOption lean::pp::notation true. +setoption lean::pp::notation true. print c |- d ; e. -Variable fact : A -> A. -Notation 20 _ ! : fact. +variable fact : A -> A. +notation 20 _ ! : fact. print c! !. -SetOption lean::pp::notation false. +setoption lean::pp::notation false. print c! !. -SetOption lean::pp::notation true. -Variable g : A -> A -> A. -Notation 30 [ _ ; _ ] : g +setoption lean::pp::notation true. +variable g : A -> A -> A. +notation 30 [ _ ; _ ] : g print [c;d]. print [c ; [d;e] ]. -SetOption lean::pp::notation false. +setoption lean::pp::notation false. print [c ; [d;e] ]. -SetOption lean::pp::notation true. -Variable h : A -> A -> A. -Notation 40 _ << _ end : h. -print Environment 1. -print d << e end. -print [c ; d << e end ]. -SetOption lean::pp::notation false. -print [c ; d << e end ]. -SetOption lean::pp::notation true. -Variable r : A -> A -> A. -Infixl 30 ++ : r. -Variable s : A -> A -> A. -Infixl 40 ** : s. +setoption lean::pp::notation true. +variable h : A -> A -> A. +notation 40 _ << _ >> : h. +print environment 1. +print d << e >>. +print [c ; d << e >> ]. +setoption lean::pp::notation false. +print [c ; d << e >> ]. +setoption lean::pp::notation true. +variable r : A -> A -> A. +infixl 30 ++ : r. +variable s : A -> A -> A. +infixl 40 ** : s. print c ** d ++ e ** c. -Variable p1 : Bool. -Variable p2 : Bool. -Variable p3 : Bool. +variable p1 : Bool. +variable p2 : Bool. +variable p3 : Bool. print p1 || p2 && p3. -SetOption lean::pp::notation false. +setoption lean::pp::notation false. print c ** d ++ e ** c. print p1 || p2 && p3. -SetOption lean::pp::notation true. +setoption lean::pp::notation true. print c = d || d = c. print not p1 || p2. print p1 && p3 || p2 && p3. -SetOption lean::pp::notation false. +setoption lean::pp::notation false. print not p1 || p2. print p1 && p3 || p2 && p3. diff --git a/tests/lean/tst3.lean.expected.out b/tests/lean/tst3.lean.expected.out index 81dbccd4d..c51f19fa0 100644 --- a/tests/lean/tst3.lean.expected.out +++ b/tests/lean/tst3.lean.expected.out @@ -1,10 +1,10 @@ Set: pp::colors Set: pp::unicode -Notation 10 if _ then _ : implies +notation 10 if _ then _ : implies if ⊤ then ⊥ if ⊤ then (if a then ⊥) implies ⊤ (implies a ⊥) -Notation 100 _ |- _ ; _ : f +notation 100 _ |- _ ; _ : f f c d e c |- d ; e (c !) ! @@ -13,9 +13,9 @@ The precedence of ';' changed from 100 to 30. [ c ; d ] [ c ; ([ d ; e ]) ] g c (g d e) -Notation 40 _ << _ end : h -d << e end -[ c ; d << e end ] +notation 40 _ << _ >> : h +d << e >> +[ c ; d << e >> ] g c (h d e) c ** d ++ e ** c p1 ∨ p2 ∧ p3 diff --git a/tests/lean/tst4.lean b/tests/lean/tst4.lean index 72f3a17d2..113b5942a 100644 --- a/tests/lean/tst4.lean +++ b/tests/lean/tst4.lean @@ -1,21 +1,21 @@ -Variable f {A : Type} (a b : A) : A -Variable N : Type -Variable n1 : N -Variable n2 : N -SetOption lean::pp::implicit true +variable f {A : Type} (a b : A) : A +variable N : Type +variable n1 : N +variable n2 : N +setoption lean::pp::implicit true print f n1 n2 print f (fun x : N -> N, x) (fun y : _, y) -Variable EqNice {A : Type} (lhs rhs : A) : Bool -Infix 50 === : EqNice +variable EqNice {A : Type} (lhs rhs : A) : Bool +infix 50 === : EqNice print n1 === n2 -Check f n1 n2 -Check @Congr +check f n1 n2 +check @Congr print f n1 n2 -Variable a : N -Variable b : N -Variable c : N -Variable g : N -> N -Axiom H1 : a = b && b = c -Theorem Pr : (g a) = (g c) := +variable a : N +variable b : N +variable c : N +variable g : N -> N +axiom H1 : a = b && b = c +theorem Pr : (g a) = (g c) := Congr (Refl g) (Trans (Conjunct1 H1) (Conjunct2 H1)) -print Environment 2 +print environment 2 diff --git a/tests/lean/tst4.lean.expected.out b/tests/lean/tst4.lean.expected.out index e6d296ec4..1ac314f31 100644 --- a/tests/lean/tst4.lean.expected.out +++ b/tests/lean/tst4.lean.expected.out @@ -18,8 +18,8 @@ Assumed: g Assumed: H1 Proved: Pr -Axiom H1 : @eq N a b ∧ @eq N b c -Theorem Pr : @eq N (g a) (g c) := +axiom H1 : @eq N a b ∧ @eq N b c +theorem Pr : @eq N (g a) (g c) := @Congr N (λ x : N, N) g diff --git a/tests/lean/tst5.lean b/tests/lean/tst5.lean index 00d8dcc49..b47f87b69 100644 --- a/tests/lean/tst5.lean +++ b/tests/lean/tst5.lean @@ -1,9 +1,9 @@ -Variable N : Type -Variable a : N -Variable b : N +variable N : Type +variable a : N +variable b : N print a = b -Check a = b -SetOption lean::pp::implicit true +check a = b +setoption lean::pp::implicit true print a = b print (Type 1) = (Type 1) print true = false diff --git a/tests/lean/tst6.lean b/tests/lean/tst6.lean index ad0a0a085..ff99e839a 100644 --- a/tests/lean/tst6.lean +++ b/tests/lean/tst6.lean @@ -1,18 +1,18 @@ -Variable N : Type -Variable h : N -> N -> N +variable N : Type +variable h : N -> N -> N -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 -- Display the theorem showing implicit arguments -SetOption lean::pp::implicit true -print Environment 2 +setoption lean::pp::implicit true +print environment 2 -- Display the theorem hiding implicit arguments -SetOption lean::pp::implicit false -print Environment 2 +setoption lean::pp::implicit false +print environment 2 -Theorem Example1 (a b c d : N) (H: (a = b ∧ b = c) ∨ (a = d ∧ d = c)) : (h a b) = (h c b) := +theorem Example1 (a b c d : N) (H: (a = b ∧ b = c) ∨ (a = d ∧ d = c)) : (h a b) = (h c b) := DisjCases H (fun H1 : a = b ∧ b = c, CongrH (Trans (Conjunct1 H1) (Conjunct2 H1)) (Refl b)) @@ -20,32 +20,32 @@ Theorem Example1 (a b c d : N) (H: (a = b ∧ b = c) ∨ (a = d ∧ d = c)) : (h CongrH (Trans (Conjunct1 H1) (Conjunct2 H1)) (Refl b)) -- print proof of the last theorem with all implicit arguments -SetOption lean::pp::implicit true -print Environment 1 +setoption lean::pp::implicit true +print environment 1 -- Using placeholders to hide the type of H1 -Theorem Example2 (a b c d : N) (H: (a = b ∧ b = c) ∨ (a = d ∧ d = c)) : (h a b) = (h c b) := +theorem Example2 (a b c d : N) (H: (a = b ∧ b = c) ∨ (a = d ∧ d = c)) : (h a b) = (h c b) := DisjCases H (fun H1 : _, CongrH (Trans (Conjunct1 H1) (Conjunct2 H1)) (Refl b)) (fun H1 : _, CongrH (Trans (Conjunct1 H1) (Conjunct2 H1)) (Refl b)) -SetOption lean::pp::implicit true -print Environment 1 +setoption lean::pp::implicit true +print environment 1 -- Same example but the first conjuct has unnecessary stuff -Theorem Example3 (a b c d e : N) (H: (a = b ∧ b = e ∧ b = c) ∨ (a = d ∧ d = c)) : (h a b) = (h c b) := +theorem Example3 (a b c d e : N) (H: (a = b ∧ b = e ∧ b = c) ∨ (a = d ∧ d = c)) : (h a b) = (h c b) := DisjCases H (fun H1 : _, CongrH (Trans (Conjunct1 H1) (Conjunct2 (Conjunct2 H1))) (Refl b)) (fun H1 : _, CongrH (Trans (Conjunct1 H1) (Conjunct2 H1)) (Refl b)) -SetOption lean::pp::implicit false -print Environment 1 +setoption lean::pp::implicit false +print environment 1 -Theorem Example4 (a b c d e : N) (H: (a = b ∧ b = e ∧ b = c) ∨ (a = d ∧ d = c)) : (h a c) = (h c a) := +theorem Example4 (a b c d e : N) (H: (a = b ∧ b = e ∧ b = c) ∨ (a = d ∧ d = c)) : (h a c) = (h c a) := DisjCases H (fun H1 : _, let AeqC := Trans (Conjunct1 H1) (Conjunct2 (Conjunct2 H1)) @@ -54,5 +54,5 @@ Theorem Example4 (a b c d e : N) (H: (a = b ∧ b = e ∧ b = c) ∨ (a = d ∧ let AeqC := Trans (Conjunct1 H1) (Conjunct2 H1) in CongrH AeqC (Symm AeqC)) -SetOption lean::pp::implicit false -print Environment 1 +setoption lean::pp::implicit false +print environment 1 diff --git a/tests/lean/tst6.lean.expected.out b/tests/lean/tst6.lean.expected.out index bdab27e94..ed6ea0e2e 100644 --- a/tests/lean/tst6.lean.expected.out +++ b/tests/lean/tst6.lean.expected.out @@ -4,15 +4,15 @@ Assumed: h Proved: CongrH Set: lean::pp::implicit -Variable h : N → N → N -Theorem CongrH {a1 a2 b1 b2 : N} (H1 : @eq N a1 b1) (H2 : @eq N a2 b2) : @eq N (h a1 a2) (h b1 b2) := +variable h : N → N → N +theorem CongrH {a1 a2 b1 b2 : N} (H1 : @eq N a1 b1) (H2 : @eq N a2 b2) : @eq N (h a1 a2) (h b1 b2) := @Congr N (λ x : N, N) (h a1) (h b1) a2 b2 (@Congr N (λ x : N, N → N) h h a1 b1 (@Refl (N → N → N) h) H1) H2 Set: lean::pp::implicit -Variable h : N → N → N -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 +variable h : N → N → N +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 Proved: Example1 Set: lean::pp::implicit -Theorem Example1 (a b c d : N) (H : @eq N a b ∧ @eq N b c ∨ @eq N a d ∧ @eq N d c) : @eq N (h a b) (h c b) := +theorem Example1 (a b c d : N) (H : @eq N a b ∧ @eq N b c ∨ @eq N a d ∧ @eq N d c) : @eq N (h a b) (h c b) := @DisjCases (@eq N a b ∧ @eq N b c) (@eq N a d ∧ @eq N d c) @@ -34,7 +34,7 @@ Theorem Example1 (a b c d : N) (H : @eq N a b ∧ @eq N b c ∨ @eq N a d ∧ @e (@Refl N b)) Proved: Example2 Set: lean::pp::implicit -Theorem Example2 (a b c d : N) (H : @eq N a b ∧ @eq N b c ∨ @eq N a d ∧ @eq N d c) : @eq N (h a b) (h c b) := +theorem Example2 (a b c d : N) (H : @eq N a b ∧ @eq N b c ∨ @eq N a d ∧ @eq N d c) : @eq N (h a b) (h c b) := @DisjCases (@eq N a b ∧ @eq N b c) (@eq N a d ∧ @eq N d c) @@ -56,14 +56,14 @@ Theorem Example2 (a b c d : N) (H : @eq N a b ∧ @eq N b c ∨ @eq N a d ∧ @e (@Refl N b)) Proved: Example3 Set: lean::pp::implicit -Theorem Example3 (a b c d e : N) (H : a = b ∧ b = e ∧ b = c ∨ a = d ∧ d = c) : h a b = h c b := +theorem Example3 (a b c d e : N) (H : a = b ∧ b = e ∧ b = c ∨ a = d ∧ d = c) : h a b = h c b := DisjCases H (λ H1 : a = b ∧ b = e ∧ b = c, CongrH (Trans (Conjunct1 H1) (Conjunct2 (Conjunct2 H1))) (Refl b)) (λ H1 : a = d ∧ d = c, CongrH (Trans (Conjunct1 H1) (Conjunct2 H1)) (Refl b)) Proved: Example4 Set: lean::pp::implicit -Theorem Example4 (a b c d e : N) (H : a = b ∧ b = e ∧ b = c ∨ a = d ∧ d = c) : h a c = h c a := +theorem Example4 (a b c d e : N) (H : a = b ∧ b = e ∧ b = c ∨ a = d ∧ d = c) : h a c = h c a := DisjCases H (λ H1 : a = b ∧ b = e ∧ b = c, diff --git a/tests/lean/tst7.lean b/tests/lean/tst7.lean index 427ffa77c..f6c6f9fe3 100644 --- a/tests/lean/tst7.lean +++ b/tests/lean/tst7.lean @@ -1,16 +1,16 @@ -Variable f : Pi (A : Type), A -> Bool +variable f : Pi (A : Type), A -> Bool print fun (A B : Type) (a : _), f B a -- The following one should produce an error print fun (A : Type) (a : _) (B : Type), f B a -Variable myeq : Pi A : (Type U), A -> A -> Bool +variable myeq : Pi A : (Type U), A -> A -> Bool print myeq _ (fun (A : Type) (a : _), a) (fun (B : Type) (b : B), b) -Check myeq _ (fun (A : Type) (a : _), a) (fun (B : Type) (b : B), b) +check myeq _ (fun (A : Type) (a : _), a) (fun (B : Type) (b : B), b) -Variable R : Type -> Bool -Variable h : (Pi (A : Type), R A) -> Bool -Check (fun (H : Bool) +variable R : Type -> Bool +variable h : (Pi (A : Type), R A) -> Bool +check (fun (H : Bool) (f1 : Pi (A : Type), R _) (g1 : Pi (A : Type), R A) (G : Pi (A : Type), myeq _ (f1 _) (g1 A)), diff --git a/tests/lean/tst8.lean b/tests/lean/tst8.lean index 8b17d278e..5340ae5c8 100644 --- a/tests/lean/tst8.lean +++ b/tests/lean/tst8.lean @@ -1,11 +1,11 @@ -Import Int -Check fun (A : Type) (a : A), +import Int +check fun (A : Type) (a : A), let b := a in b -Variable g : Pi A : Type, A -> A +variable g : Pi A : Type, A -> A -Definition f (A: Type) (a : A) : A := +definition f (A: Type) (a : A) : A := let b := g A a, c := g A b in c diff --git a/tests/lean/tst9.lean b/tests/lean/tst9.lean index 89f5efc9b..3bc77d7f0 100644 --- a/tests/lean/tst9.lean +++ b/tests/lean/tst9.lean @@ -1,5 +1,5 @@ -Variable f : Pi A : Type, A -> A -> A -Variable N : Type -Variable g : N -> N -> Bool -Variable a : N -Check g true (f _ a a) +variable f : Pi A : Type, A -> A -> A +variable N : Type +variable g : N -> N -> Bool +variable a : N +check g true (f _ a a) diff --git a/tests/lean/ty1.lean b/tests/lean/ty1.lean index da105c81a..8d086a9d4 100644 --- a/tests/lean/ty1.lean +++ b/tests/lean/ty1.lean @@ -1,9 +1,9 @@ -Import Int. -Variable i : Int. -Check fun x, x + i -Check fun x, x + 1 -Check fun x, x -Check fun x y, y + i + 1 + x -Check (fun x, x) i -Check (fun x, x i) (fun x y, x + 1 + y) -Check (fun x, x) (fun x y, x + 1 + y) +import Int. +variable i : Int. +check fun x, x + i +check fun x, x + 1 +check fun x, x +check fun x y, y + i + 1 + x +check (fun x, x) i +check (fun x, x i) (fun x y, x + 1 + y) +check (fun x, x) (fun x y, x + 1 + y) diff --git a/tests/lean/ty2.lean b/tests/lean/ty2.lean index 31d897482..da1d2c007 100644 --- a/tests/lean/ty2.lean +++ b/tests/lean/ty2.lean @@ -1,6 +1,6 @@ -Definition B : Type := Bool -Definition T : (Type 1) := Type -Variable N : T -Variable x : N -Variable a : B -Axiom H : a +definition B : Type := Bool +definition T : (Type 1) := Type +variable N : T +variable x : N +variable a : B +axiom H : a diff --git a/tests/lean/type_inf_bug1.lean b/tests/lean/type_inf_bug1.lean index c30909be9..5145d4b02 100644 --- a/tests/lean/type_inf_bug1.lean +++ b/tests/lean/type_inf_bug1.lean @@ -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) (a : A) (b : A') (L2 : A' == A), @@ -9,7 +9,7 @@ Check fun (A A': TypeM) L3 : b == b' := CastEq L2 b in L3 -Check fun (A A': TypeM) +check fun (A A': TypeM) (B : A -> TypeM) (B' : A' -> TypeM) (f : Pi x : A, B x) diff --git a/tests/lean/unicode.lean b/tests/lean/unicode.lean index 33a56f089..bcec1d33a 100644 --- a/tests/lean/unicode.lean +++ b/tests/lean/unicode.lean @@ -1,8 +1,8 @@ print true /\ false -SetOption pp::unicode false +setoption pp::unicode false print true /\ false -SetOption pp::unicode true -SetOption lean::pp::notation false +setoption pp::unicode true +setoption lean::pp::notation false print true /\ false -SetOption pp::unicode false +setoption pp::unicode false print true /\ false diff --git a/tests/lean/vars1.lean b/tests/lean/vars1.lean index 91e36ffe6..08a17d0e9 100644 --- a/tests/lean/vars1.lean +++ b/tests/lean/vars1.lean @@ -1,5 +1,5 @@ -Import Int. -Variables a b c : Int -Variables d b e : Int -Variables d e f : Int -Eval 1 + 1 + a + b + c + d + e + f \ No newline at end of file +import Int. +variables a b c : Int +variables d b e : Int +variables d e f : Int +eval 1 + 1 + a + b + c + d + e + f \ No newline at end of file diff --git a/tests/lua/coercion_bug1.lua b/tests/lua/coercion_bug1.lua index 5b442ad49..ef89a4b49 100644 --- a/tests/lua/coercion_bug1.lua +++ b/tests/lua/coercion_bug1.lua @@ -1,10 +1,10 @@ local env = environment() env:import("Int") parse_lean_cmds([[ - Variable f : Int -> Int -> Int - Notation 20 _ +++ _ : f + variable f : Int -> Int -> Int + notation 20 _ +++ _ : f print f 10 20 - Notation 20 _ -+- _ : f + notation 20 _ -+- _ : f print f 10 20 ]], env) diff --git a/tests/lua/front.lua b/tests/lua/front.lua index 42a214424..157576be5 100644 --- a/tests/lua/front.lua +++ b/tests/lua/front.lua @@ -3,15 +3,15 @@ local env = environment() env:import("Int") print(get_options()) parse_lean_cmds([[ -Variable f : Int -> Int -> Int -Variable g : Bool -> Bool -> Bool -Variables a b : Int -Variables i j : Int -Variables p q : Bool -Notation 100 _ ++ _ : f -Notation 100 _ ++ _ : g -SetOption pp::colors true -SetOption pp::width 300 +variable f : Int -> Int -> Int +variable g : Bool -> Bool -> Bool +variables a b : Int +variables i j : Int +variables p q : Bool +notation 100 _ ++ _ : f +notation 100 _ ++ _ : g +setoption pp::colors true +setoption pp::width 300 ]], env) print(get_options()) assert(get_options():get{"pp", "colors"}) @@ -24,10 +24,10 @@ parse_lean_cmds([[ local env2 = environment() env2:import("Int") parse_lean_cmds([[ -Variable f : Int -> Int -> Int -Variables a b : Int +variable f : Int -> Int -> Int +variables a b : Int print f a b -Notation 100 _ -+ _ : f +notation 100 _ -+ _ : f ]], env2) local f, a, b = Consts("f, a, b") diff --git a/tests/lua/hidden1.lua b/tests/lua/hidden1.lua index d5cfefad4..3db1650de 100644 --- a/tests/lua/hidden1.lua +++ b/tests/lua/hidden1.lua @@ -4,7 +4,7 @@ assert(env:is_opaque("and")) assert(env:is_opaque("or")) assert(env:is_opaque({"Int", "lt"})) parse_lean_cmds([[ - Definition a : Bool := true + definition a : Bool := true ]], env) assert(not env:is_opaque("a")) env:set_opaque("a", true) diff --git a/tests/lua/parser2.lua b/tests/lua/parser2.lua index e9740678c..fd6389ddd 100644 --- a/tests/lua/parser2.lua +++ b/tests/lua/parser2.lua @@ -1,26 +1,26 @@ import("util.lua") local env = environment() parse_lean_cmds([[ - Variable N : Type - Variables x y : N - Variable f : N -> N -> N - SetOption pp::colors false + variable N : Type + variables x y : N + variable f : N -> N -> N + setoption pp::colors false ]], env) local f, x, y = Consts("f, x, y") print(env:type_check(f(x, y))) assert(env:type_check(f(x, y)) == Const("N")) assert(not get_options():get{"pp", "colors"}) parse_lean_cmds([[ - SetOption pp::colors true + setoption pp::colors true ]], env) assert(get_options():get{"pp", "colors"}) local o = get_options() o:update({"lean", "pp", "notation"}, false) assert(not o:get{"lean", "pp", "notation"}) o = parse_lean_cmds([[ - Check fun x : N, y - SetOption pp::notation true - Check fun x : N, y + check fun x : N, y + setoption pp::notation true + check fun x : N, y ]], env, o) print(o) assert(o:get{"lean", "pp", "notation"}) diff --git a/tests/lua/single.lua b/tests/lua/single.lua index ca6fbc1d2..204245722 100644 --- a/tests/lua/single.lua +++ b/tests/lua/single.lua @@ -1,5 +1,5 @@ print("hello world") local env = environment() env:import("Int") -parse_lean_cmds([[ Variables a b : Int ]], env) +parse_lean_cmds([[ variables a b : Int ]], env) print(parse_lean([[a + b + 10]], env)) diff --git a/tests/lua/template1.lua b/tests/lua/template1.lua index 67dd9f95c..f3e79dfa7 100644 --- a/tests/lua/template1.lua +++ b/tests/lua/template1.lua @@ -3,8 +3,8 @@ import("template.lua") local env = environment() env:import("Int") parse_lean_cmds([[ - Variables a b c : Int - Variables f : Int -> Int + variables a b c : Int + variables f : Int -> Int ]], env) local a, b, c = Consts("a, b, c") print(parse_template("%1 + f %2 + %1 + %1", {a, b}, env))