chore(*): minimize dependencies on tests

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
Leonardo de Moura 2014-08-24 19:58:48 -07:00
parent fbf13994d8
commit cbc81ea6c5
84 changed files with 84 additions and 107 deletions

View file

@ -1,4 +1,4 @@
import standard import logic
using bool eq_ops tactic using bool eq_ops tactic
variables a b c : bool variables a b c : bool

View file

@ -1,4 +1,4 @@
import standard import logic
namespace tst namespace tst
definition foo {A B : Type} (a : A) (b : B) := a definition foo {A B : Type} (a : A) (b : B) := a

View file

@ -1,4 +1,4 @@
import standard import logic
using num using num
abbreviation Type1 := Type.{1} abbreviation Type1 := Type.{1}

View file

@ -1,4 +1,4 @@
import standard import data.bool
using bool using bool
check ff check ff

View file

@ -1,4 +1,4 @@
import standard import logic
theorem symm2 {A : Type} {a b : A} (H : a = b) : b = a theorem symm2 {A : Type} {a b : A} (H : a = b) : b = a
:= subst H (refl a) := subst H (refl a)

View file

@ -1,4 +1,4 @@
import standard import logic
section section
parameter {A : Type} parameter {A : Type}
theorem T {a b : A} (H : a = b) : b = a theorem T {a b : A} (H : a = b) : b = a

View file

@ -1,4 +1,4 @@
import standard import data.num
using num using num
namespace foo namespace foo

View file

@ -1,4 +1,4 @@
import standard import logic data.prod
using num prod inhabited using num prod inhabited
definition H : inhabited (Prop × num × (num → num)) := _ definition H : inhabited (Prop × num × (num → num)) := _

View file

@ -1,4 +1,4 @@
import standard import logic data.prod
using num prod nonempty inhabited using num prod nonempty inhabited
theorem H {A B : Type} (H1 : inhabited A) : inhabited (Prop × A × (B → num)) theorem H {A B : Type} (H1 : inhabited A) : inhabited (Prop × A × (B → num))

View file

@ -1,4 +1,4 @@
import standard import logic data.prod
using num prod inhabited using num prod inhabited
section section

View file

@ -1,4 +1,4 @@
import standard import logic
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,4 +1,4 @@
import standard import logic
namespace algebra namespace algebra
inductive mul_struct (A : Type) : Type := inductive mul_struct (A : Type) : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic data.prod
using prod using prod
inductive t1 : Type := inductive t1 : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using num tactic using num tactic
inductive inh (A : Type) : Type := inductive inh (A : Type) : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic data.prod
using num tactic prod using num tactic prod
inductive inh (A : Type) : Prop := inductive inh (A : Type) : Prop :=

View file

@ -1,4 +1,4 @@
import standard import logic
namespace setoid namespace setoid
inductive setoid : Type := inductive setoid : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
namespace setoid namespace setoid
inductive setoid : Type := inductive setoid : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
namespace setoid namespace setoid
inductive setoid : Type := inductive setoid : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
namespace setoid namespace setoid
inductive setoid : Type := inductive setoid : Type :=

View file

@ -1,4 +1,4 @@
import standard data.unit import logic data.unit
using bool unit decidable using bool unit decidable
variables a b c : bool variables a b c : bool

View file

@ -4,7 +4,7 @@
--- Author: Jeremy Avigad --- Author: Jeremy Avigad
---------------------------------------------------------------------------------------------------- ----------------------------------------------------------------------------------------------------
import standard struc.function import logic struc.function
using function using function

View file

@ -1,4 +1,4 @@
import standard import logic
using num using num
variable p : num → num → num → Prop variable p : num → num → num → Prop
axiom H1 : ∃ x y z, p x y z axiom H1 : ∃ x y z, p x y z

View file

@ -1,4 +1,4 @@
import standard import logic
using num tactic using num tactic
variable p : num → num → num → Prop variable p : num → num → num → Prop
axiom H1 : ∃ x y z, p x y z axiom H1 : ∃ x y z, p x y z

View file

@ -1,4 +1,4 @@
import standard import logic
namespace foo namespace foo
variable x : num.num variable x : num.num
check x check x

View file

@ -1,4 +1,4 @@
import standard struc.function import logic struc.function
using function num bool using function num bool

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem T {a b c d : Prop} (H : a) (H : b) (H : c) (H : d) : a theorem T {a b c d : Prop} (H : a) (H : b) (H : c) (H : d) : a

View file

@ -1,4 +1,4 @@
import standard import logic
using num using num
section section

View file

@ -1,4 +1,4 @@
import standard import logic
using num using num
section section

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
variables a b c d : Prop variables a b c d : Prop
axiom Ha : a axiom Ha : a

View file

@ -1,4 +1,4 @@
import standard import logic
definition id {A : Type} (a : A) := a definition id {A : Type} (a : A) := a
check id id check id id
set_option pp.universes true set_option pp.universes true

View file

@ -1,4 +1,4 @@
import standard import logic
definition f {A : Type} {B : Type} (f : A → B → Prop) ⦃C : Type⦄ {R : C → C → Prop} {c : C} (H : R c c) : R c c definition f {A : Type} {B : Type} (f : A → B → Prop) ⦃C : Type⦄ {R : C → C → Prop} {c : C} (H : R c c) : R c c
:= H := H
@ -7,4 +7,3 @@ variable g : Prop → Prop → Prop
variable H : true ∧ true variable H : true ∧ true
check f g H check f g H

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
inductive list (A : Type) : Type := inductive list (A : Type) : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
check check
let f x y := x ∧ y, let f x y := x ∧ y,

View file

@ -1,4 +1,4 @@
import standard import logic
using num eq_ops using num eq_ops
inductive nat : Type := inductive nat : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using num eq_ops using num eq_ops
inductive nat : Type := inductive nat : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using num eq_ops using num eq_ops
inductive nat : Type := inductive nat : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using num eq_ops using num eq_ops
inductive nat : Type := inductive nat : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using eq_ops using eq_ops
inductive nat : Type := inductive nat : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using bool using bool
variable list : Type.{1} variable list : Type.{1}
@ -19,4 +19,3 @@ check a :: b :: nil
check [a, b] check [a, b]
check [a, b, c] check [a, b, c]
check [] check []

View file

@ -1,4 +1,4 @@
import standard import logic
namespace foo namespace foo
namespace boo namespace boo

View file

@ -1,4 +1,4 @@
import standard import logic
check 14 check 14
check 0 check 0
check 3 check 3

View file

@ -1,4 +1,4 @@
import standard import logic data.prod
using prod using prod
-- Test tuple notation -- Test tuple notation

View file

@ -1,4 +1,4 @@
import standard import logic
using num using num
variable foo : Prop variable foo : Prop
@ -15,6 +15,3 @@ namespace N1
print raw _root_.foo print raw _root_.foo
end N2 end N2
end N1 end N1

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
section section

View file

@ -1,4 +1,4 @@
import standard import logic
using bool using bool
definition set {{T : Type}} := T → bool definition set {{T : Type}} := T → bool

View file

@ -1,4 +1,4 @@
import standard import logic
using bool using bool
namespace set namespace set

View file

@ -1,3 +1,3 @@
import standard import logic
using tactic using tactic
print raw (by assumption) print raw (by assumption)

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : A theorem tst {A B : Prop} (H1 : A) (H2 : B) : A

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst (a b : Prop) (H : ¬ a ¬ b) (Hb : b) : ¬ a ∧ b theorem tst (a b : Prop) (H : ¬ a ¬ b) (Hb : b) : ¬ a ∧ b
@ -9,8 +9,3 @@ theorem tst (a b : Prop) (H : ¬ a ¬ b) (Hb : b) : ¬ a ∧ b
(assume Hna, absurd Ha Hna) (assume Hna, absurd Ha Hna)
(assume Hnb, absurd Hb Hnb)); (assume Hnb, absurd Hb Hnb));
assumption assumption

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst (a b : Prop) (H : ¬ a ¬ b) (Hb : b) : ¬ a ∧ b := theorem tst (a b : Prop) (H : ¬ a ¬ b) (Hb : b) : ¬ a ∧ b :=

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
definition basic_tac : tactic definition basic_tac : tactic

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
variable A : Type.{1} variable A : Type.{1}

View file

@ -1,4 +1,4 @@
import standard data.string import logic data.string
using tactic using tactic
variable A : Type.{1} variable A : Type.{1}

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
variable A : Type.{1} variable A : Type.{1}
@ -8,5 +8,3 @@ theorem tst {a b c : A} (H1 : a = b) (H2 : b = c) : f a b = f b c
:= by apply (@congr A A (f a) (f b)); := by apply (@congr A A (f a) (f b));
apply (congr_arg f); apply (congr_arg f);
!assumption !assumption

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
variable A : Type.{1} variable A : Type.{1}
@ -9,4 +9,3 @@ theorem tst {a b c : A} (H1 : a = b) (H2 : b = c) : f a b = f b c
apply (subst H2); apply (subst H2);
apply refl; apply refl;
assumption assumption

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst {A : Type} {f : A → A → A} {a b c : A} (H1 : a = b) (H2 : b = c) : f a b = f b c theorem tst {A : Type} {f : A → A → A} {a b c : A} (H1 : a = b) (H2 : b = c) : f a b = f b c
@ -6,4 +6,3 @@ theorem tst {A : Type} {f : A → A → A} {a b c : A} (H1 : a = b) (H2 : b = c)
apply (subst H2); apply (subst H2);
apply refl; apply refl;
assumption assumption

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : A theorem tst {A B : Prop} (H1 : A) (H2 : B) : A

View file

@ -1,12 +1,7 @@
import standard import logic
using tactic using tactic
definition assump := eassumption definition assump := eassumption
theorem tst {A : Type} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c theorem tst {A : Type} {a b c : A} (H1 : a = b) (H2 : b = c) : a = c
:= by apply trans; assump; assump := by apply trans; assump; assump

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
definition assump := eassumption definition assump := eassumption
@ -12,6 +12,3 @@ theorem tst2 {A : Type} {a b c d : A} {p : A → A → Prop} (Ha : p a c) (H1 :
(* (*
print(get_env():find("tst2"):value()) print(get_env():find("tst2"):value())
*) *)

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem T (a b c d : Prop) (Ha : a) (Hb : b) (Hc : c) (Hd : d) : a ∧ b ∧ c ∧ d theorem T (a b c d : Prop) (Ha : a) (Hb : b) (Hc : c) (Hd : d) : a ∧ b ∧ c ∧ d

View file

@ -1,4 +1,4 @@
import standard import logic
using num (num pos_num num_rec pos_num_rec) using num (num pos_num num_rec pos_num_rec)
using tactic using tactic

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
definition my_tac1 := apply @refl definition my_tac1 := apply @refl

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
definition my_tac1 := apply @refl definition my_tac1 := apply @refl

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic inhabited using tactic inhabited
inductive sum (A : Type) (B : Type) : Type := inductive sum (A : Type) (B : Type) : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
definition my_tac := repeat ([ apply @and_intro definition my_tac := repeat ([ apply @and_intro

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic inhabited using tactic inhabited
inductive sum (A : Type) (B : Type) : Type := inductive sum (A : Type) (B : Type) : Type :=

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
section section

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : A theorem tst {A B : Prop} (H1 : A) (H2 : B) : A

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
section section

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic (renaming id->id_tac) using tactic (renaming id->id_tac)
definition id {A : Type} (a : A) := a definition id {A : Type} (a : A) := a

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic (renaming id->id_tac) using tactic (renaming id->id_tac)
definition id {A : Type} (a : A) := a definition id {A : Type} (a : A) := a

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic (renaming id->id_tac) using tactic (renaming id->id_tac)
definition id {A : Type} (a : A) := a definition id {A : Type} (a : A) := a

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A theorem tst {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A

View file

@ -1,4 +1,4 @@
import standard import logic
using tactic using tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A theorem tst {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A

View file

@ -1,6 +1,5 @@
import standard import logic
using tactic using tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : ((fun x : Prop, x) A) ∧ B ∧ A theorem tst {A B : Prop} (H1 : A) (H2 : B) : ((fun x : Prop, x) A) ∧ B ∧ A
:= by apply and_intro; beta; assumption; apply and_intro; !assumption := by apply and_intro; beta; assumption; apply and_intro; !assumption

View file

@ -1,4 +1,4 @@
import standard import logic
using num using num
definition proj1 (x : num) (y : num) := x definition proj1 (x : num) (y : num) := x

View file

@ -1,4 +1,4 @@
import standard import logic
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,4 +1,4 @@
import standard import logic
variable N : Type variable N : Type
variable α : N variable α : N

View file

@ -1,4 +1,4 @@
import standard import logic
namespace S1 namespace S1
hypothesis I : Type hypothesis I : Type

View file

@ -1,4 +1,4 @@
import standard import logic
hypothesis I : Type hypothesis I : Type
definition F (X : Type) : Type := (X → Prop) → Prop definition F (X : Type) : Type := (X → Prop) → Prop

View file

@ -1,4 +1,4 @@
import standard import logic
using bool eq_ops tactic using bool eq_ops tactic
variables a b c : bool variables a b c : bool

View file

@ -3,7 +3,7 @@
-- Released under Apache 2.0 license as described in the file LICENSE. -- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Floris van Doorn -- Author: Floris van Doorn
---------------------------------------------------------------------------------------------------- ----------------------------------------------------------------------------------------------------
import standard struc.binary import logic struc.binary
using tactic num binary eq_ops using tactic num binary eq_ops
using decidable using decidable