feat(frontends/lean): rename '[intro]' ==> '[intro!]' and '[backward]' ==> '[intro]'

This commit is contained in:
Leonardo de Moura 2015-12-07 21:33:35 -08:00
parent 041c1cbb17
commit 50df6b5698
16 changed files with 56 additions and 60 deletions

View file

@ -56,7 +56,7 @@ section
... = f₂ : rfl ... = f₂ : rfl
end end
attribute funext [intro] attribute funext [intro!]
open function.equiv_notation open function.equiv_notation

View file

@ -24,7 +24,7 @@ prefix `¬` := not
definition absurd {a : Prop} {b : Type} (H1 : a) (H2 : ¬a) : b := definition absurd {a : Prop} {b : Type} (H1 : a) (H2 : ¬a) : b :=
false.rec b (H2 H1) false.rec b (H2 H1)
lemma not.intro [intro] {a : Prop} (H : a → false) : ¬ a := lemma not.intro [intro!] {a : Prop} (H : a → false) : ¬ a :=
H H
theorem mt {a b : Prop} (H1 : a → b) (H2 : ¬b) : ¬a := theorem mt {a b : Prop} (H1 : a → b) (H2 : ¬b) : ¬a :=
@ -243,7 +243,7 @@ notation a ∧ b := and a b
variables {a b c d : Prop} variables {a b c d : Prop}
attribute and.rec [elim] attribute and.rec [elim]
attribute and.intro [intro] attribute and.intro [intro!]
theorem and.elim (H₁ : a ∧ b) (H₂ : a → b → c) : c := theorem and.elim (H₁ : a ∧ b) (H₂ : a → b → c) : c :=
and.rec H₂ H₁ and.rec H₂ H₁
@ -280,7 +280,7 @@ notation a ↔ b := iff a b
theorem iff.intro : (a → b) → (b → a) → (a ↔ b) := and.intro theorem iff.intro : (a → b) → (b → a) → (a ↔ b) := and.intro
attribute iff.intro [intro] attribute iff.intro [intro!]
theorem iff.elim : ((a → b) → (b → a) → c) → (a ↔ b) → c := and.rec theorem iff.elim : ((a → b) → (b → a) → c) → (a ↔ b) → c := and.rec
@ -508,9 +508,7 @@ and_congr (imp_congr H1 H2) (imp_congr H2 H1)
inductive Exists {A : Type} (P : A → Prop) : Prop := inductive Exists {A : Type} (P : A → Prop) : Prop :=
intro : ∀ (a : A), P a → Exists P intro : ∀ (a : A), P a → Exists P
-- Remark: we don't mark exists.intro with [intro] because the [intro] attribute Exists.intro [intro]
-- rules are applied eagerly.
attribute Exists.intro [backward]
definition exists.intro := @Exists.intro definition exists.intro := @Exists.intro
@ -530,7 +528,7 @@ definition exists_unique {A : Type} (p : A → Prop) :=
notation `∃!` binders `, ` r:(scoped P, exists_unique P) := r notation `∃!` binders `, ` r:(scoped P, exists_unique P) := r
theorem exists_unique.intro [backward] {A : Type} {p : A → Prop} (w : A) (H1 : p w) (H2 : ∀y, p y → y = w) : theorem exists_unique.intro [intro] {A : Type} {p : A → Prop} (w : A) (H1 : p w) (H2 : ∀y, p y → y = w) :
∃!x, p x := ∃!x, p x :=
exists.intro w (and.intro H1 H2) exists.intro w (and.intro H1 H2)

View file

@ -55,9 +55,9 @@
"class" "parsing_only" "coercion" "unfold_full" "constructor" "class" "parsing_only" "coercion" "unfold_full" "constructor"
"reducible" "irreducible" "semireducible" "quasireducible" "wf" "reducible" "irreducible" "semireducible" "quasireducible" "wf"
"whnf" "multiple_instances" "none" "decls" "declarations" "whnf" "multiple_instances" "none" "decls" "declarations"
"coercions" "classes" "symm" "subst" "refl" "trans" "simp" "simps" "congr" "backward" "coercions" "classes" "symm" "subst" "refl" "trans" "simp" "simps" "congr"
"forward" "no_pattern" "notations" "abbreviations" "begin_end_hints" "tactic_hints" "forward" "no_pattern" "notations" "abbreviations" "begin_end_hints" "tactic_hints"
"reduce_hints" "unfold_hints" "aliases" "eqv" "intro" "elim" "reduce_hints" "unfold_hints" "aliases" "eqv" "intro" "intro!" "elim"
"localrefinfo" "recursor")) "localrefinfo" "recursor"))
"lean modifiers") "lean modifiers")
(defconst lean-modifiers-regexp (defconst lean-modifiers-regexp

View file

@ -709,7 +709,7 @@ environment print_cmd(parser & p) {
} else if (p.curr_is_token(get_simp_attr_tk())) { } else if (p.curr_is_token(get_simp_attr_tk())) {
p.next(); p.next();
print_simp_rules(p); print_simp_rules(p);
} else if (p.curr_is_token(get_intro_attr_tk())) { } else if (p.curr_is_token(get_intro_bang_attr_tk())) {
p.next(); p.next();
print_intro_lemmas(p); print_intro_lemmas(p);
} else if (p.curr_is_token(get_elim_attr_tk())) { } else if (p.curr_is_token(get_elim_attr_tk())) {
@ -722,7 +722,7 @@ environment print_cmd(parser & p) {
p.next(); p.next();
p.check_token_next(get_rbracket_tk(), "invalid 'print [light]', ']' expected"); p.check_token_next(get_rbracket_tk(), "invalid 'print [light]', ']' expected");
print_light_rules(p); print_light_rules(p);
} else if (p.curr_is_token(get_backward_attr_tk())) { } else if (p.curr_is_token(get_intro_attr_tk())) {
p.next(); p.next();
print_backward_rules(p); print_backward_rules(p);
} else if (print_polymorphic(p)) { } else if (print_polymorphic(p)) {

View file

@ -46,9 +46,9 @@ decl_attributes::decl_attributes(bool is_abbrev, bool persistent):
m_recursor = false; m_recursor = false;
m_simp = false; m_simp = false;
m_congr = false; m_congr = false;
m_backward = false;
m_forward = false; m_forward = false;
m_intro = false; m_intro = false;
m_intro_bang = false;
m_elim = false; m_elim = false;
m_no_pattern = false; m_no_pattern = false;
} }
@ -97,9 +97,9 @@ void decl_attributes::parse(parser & p) {
p.next(); p.next();
} else if (auto it = parse_priority(p)) { } else if (auto it = parse_priority(p)) {
m_priority = *it; m_priority = *it;
if (!m_is_instance && !m_simp && !m_congr && !m_backward && !m_forward && !m_elim && !m_intro) { if (!m_is_instance && !m_simp && !m_congr && !m_forward && !m_elim && !m_intro && !m_intro_bang) {
throw parser_error("invalid '[priority]' attribute, declaration must be marked as an '[instance]', '[simp]', '[congr]', " throw parser_error("invalid '[priority]' attribute, declaration must be marked as an '[instance]', '[simp]', '[congr]', "
"'[backward]', '[forward]', '[intro]' or '[elim]'", pos); "'[forward]', '[intro!]', '[intro]' or '[elim]'", pos);
} }
} else if (p.curr_is_token(get_parsing_only_tk())) { } else if (p.curr_is_token(get_parsing_only_tk())) {
if (!m_is_abbrev) if (!m_is_abbrev)
@ -148,9 +148,6 @@ void decl_attributes::parse(parser & p) {
p.next(); p.next();
m_light_arg = p.parse_small_nat(); m_light_arg = p.parse_small_nat();
p.check_token_next(get_rbracket_tk(), "light attribute has form '[light <i>]'"); p.check_token_next(get_rbracket_tk(), "light attribute has form '[light <i>]'");
} else if (p.curr_is_token(get_backward_attr_tk())) {
p.next();
m_backward = true;
} else if (p.curr_is_token(get_forward_attr_tk())) { } else if (p.curr_is_token(get_forward_attr_tk())) {
p.next(); p.next();
m_forward = true; m_forward = true;
@ -160,6 +157,9 @@ void decl_attributes::parse(parser & p) {
} else if (p.curr_is_token(get_intro_attr_tk())) { } else if (p.curr_is_token(get_intro_attr_tk())) {
p.next(); p.next();
m_intro = true; m_intro = true;
} else if (p.curr_is_token(get_intro_bang_attr_tk())) {
p.next();
m_intro_bang = true;
} else if (p.curr_is_token(get_no_pattern_attr_tk())) { } else if (p.curr_is_token(get_no_pattern_attr_tk())) {
p.next(); p.next();
m_no_pattern = true; m_no_pattern = true;
@ -252,12 +252,6 @@ environment decl_attributes::apply(environment env, io_state const & ios, name c
if (m_light_arg) { if (m_light_arg) {
env = add_light_rule(env, d, *m_light_arg, ns, m_persistent); env = add_light_rule(env, d, *m_light_arg, ns, m_persistent);
} }
if (m_backward) {
if (m_priority)
env = add_backward_rule(env, d, *m_priority, ns, m_persistent);
else
env = add_backward_rule(env, d, LEAN_BACKWARD_DEFAULT_PRIORITY, ns, m_persistent);
}
if (forward) { if (forward) {
mk_multipatterns(env, ios, d); // try to create patterns mk_multipatterns(env, ios, d); // try to create patterns
if (m_priority) if (m_priority)
@ -266,6 +260,12 @@ environment decl_attributes::apply(environment env, io_state const & ios, name c
env = add_forward_lemma(env, d, LEAN_FORWARD_DEFAULT_PRIORITY, ns, m_persistent); env = add_forward_lemma(env, d, LEAN_FORWARD_DEFAULT_PRIORITY, ns, m_persistent);
} }
if (m_intro) { if (m_intro) {
if (m_priority)
env = add_backward_rule(env, d, *m_priority, ns, m_persistent);
else
env = add_backward_rule(env, d, LEAN_BACKWARD_DEFAULT_PRIORITY, ns, m_persistent);
}
if (m_intro_bang) {
if (m_priority) if (m_priority)
env = add_intro_lemma(env, ios, d, *m_priority, ns, m_persistent); env = add_intro_lemma(env, ios, d, *m_priority, ns, m_persistent);
else else
@ -290,8 +290,8 @@ void decl_attributes::write(serializer & s) const {
<< m_is_reducible << m_is_irreducible << m_is_semireducible << m_is_quasireducible << m_is_reducible << m_is_irreducible << m_is_semireducible << m_is_quasireducible
<< m_is_class << m_is_parsing_only << m_has_multiple_instances << m_unfold_full_hint << m_is_class << m_is_parsing_only << m_has_multiple_instances << m_unfold_full_hint
<< m_constructor_hint << m_symm << m_trans << m_refl << m_subst << m_recursor << m_constructor_hint << m_symm << m_trans << m_refl << m_subst << m_recursor
<< m_simp << m_congr << m_recursor_major_pos << m_priority << m_backward << m_simp << m_congr << m_recursor_major_pos << m_priority
<< m_forward << m_elim << m_intro << m_no_pattern; << m_forward << m_elim << m_intro << m_intro_bang << m_no_pattern;
write_list(s, m_unfold_hint); write_list(s, m_unfold_hint);
} }
@ -300,8 +300,8 @@ void decl_attributes::read(deserializer & d) {
>> m_is_reducible >> m_is_irreducible >> m_is_semireducible >> m_is_quasireducible >> m_is_reducible >> m_is_irreducible >> m_is_semireducible >> m_is_quasireducible
>> m_is_class >> m_is_parsing_only >> m_has_multiple_instances >> m_unfold_full_hint >> m_is_class >> m_is_parsing_only >> m_has_multiple_instances >> m_unfold_full_hint
>> m_constructor_hint >> m_symm >> m_trans >> m_refl >> m_subst >> m_recursor >> m_constructor_hint >> m_symm >> m_trans >> m_refl >> m_subst >> m_recursor
>> m_simp >> m_congr >> m_recursor_major_pos >> m_priority >> m_backward >> m_simp >> m_congr >> m_recursor_major_pos >> m_priority
>> m_forward >> m_elim >> m_intro >> m_no_pattern; >> m_forward >> m_elim >> m_intro >> m_intro_bang >> m_no_pattern;
m_unfold_hint = read_list<unsigned>(d); m_unfold_hint = read_list<unsigned>(d);
} }
} }

View file

@ -31,9 +31,9 @@ class decl_attributes {
bool m_recursor; bool m_recursor;
bool m_simp; bool m_simp;
bool m_congr; bool m_congr;
bool m_backward;
bool m_forward; bool m_forward;
bool m_intro; bool m_intro;
bool m_intro_bang;
bool m_elim; bool m_elim;
bool m_no_pattern; bool m_no_pattern;
optional<unsigned> m_recursor_major_pos; optional<unsigned> m_recursor_major_pos;

View file

@ -111,8 +111,8 @@ void init_token_table(token_table & t) {
"[persistent]", "[visible]", "[instance]", "[trans_instance]", "[light", "[persistent]", "[visible]", "[instance]", "[trans_instance]", "[light",
"[none]", "[class]", "[coercion]", "[reducible]", "[irreducible]", "[semireducible]", "[quasireducible]", "[none]", "[class]", "[coercion]", "[reducible]", "[irreducible]", "[semireducible]", "[quasireducible]",
"[simp]", "[congr]", "[parsing_only]", "[multiple_instances]", "[symm]", "[trans]", "[refl]", "[subst]", "[recursor", "[simp]", "[congr]", "[parsing_only]", "[multiple_instances]", "[symm]", "[trans]", "[refl]", "[subst]", "[recursor",
"evaluate", "check", "eval", "[wf]", "[whnf]", "[priority", "[unfold_full]", "[unfold_hints]", "[backward]", "evaluate", "check", "eval", "[wf]", "[whnf]", "[priority", "[unfold_full]", "[unfold_hints]", "[intro]",
"[forward]", "[no_pattern]", "[constructor]", "[elim]", "[intro]", "[unfold", "print", "[forward]", "[no_pattern]", "[constructor]", "[elim]", "[intro!]", "[unfold", "print",
"end", "namespace", "section", "prelude", "help", "end", "namespace", "section", "prelude", "help",
"import", "inductive", "record", "structure", "module", "universe", "universes", "local", "import", "inductive", "record", "structure", "module", "universe", "universes", "local",
"precedence", "reserve", "infixl", "infixr", "infix", "postfix", "prefix", "notation", "precedence", "reserve", "infixl", "infixr", "infix", "postfix", "prefix", "notation",

View file

@ -121,10 +121,10 @@ static name const * g_subst_tk = nullptr;
static name const * g_simp_attr_tk = nullptr; static name const * g_simp_attr_tk = nullptr;
static name const * g_congr_attr_tk = nullptr; static name const * g_congr_attr_tk = nullptr;
static name const * g_light_attr_tk = nullptr; static name const * g_light_attr_tk = nullptr;
static name const * g_backward_attr_tk = nullptr;
static name const * g_no_pattern_attr_tk = nullptr; static name const * g_no_pattern_attr_tk = nullptr;
static name const * g_forward_attr_tk = nullptr; static name const * g_forward_attr_tk = nullptr;
static name const * g_intro_attr_tk = nullptr; static name const * g_intro_attr_tk = nullptr;
static name const * g_intro_bang_attr_tk = nullptr;
static name const * g_elim_attr_tk = nullptr; static name const * g_elim_attr_tk = nullptr;
static name const * g_recursor_tk = nullptr; static name const * g_recursor_tk = nullptr;
static name const * g_attribute_tk = nullptr; static name const * g_attribute_tk = nullptr;
@ -281,10 +281,10 @@ void initialize_tokens() {
g_simp_attr_tk = new name{"[simp]"}; g_simp_attr_tk = new name{"[simp]"};
g_congr_attr_tk = new name{"[congr]"}; g_congr_attr_tk = new name{"[congr]"};
g_light_attr_tk = new name{"[light"}; g_light_attr_tk = new name{"[light"};
g_backward_attr_tk = new name{"[backward]"};
g_no_pattern_attr_tk = new name{"[no_pattern]"}; g_no_pattern_attr_tk = new name{"[no_pattern]"};
g_forward_attr_tk = new name{"[forward]"}; g_forward_attr_tk = new name{"[forward]"};
g_intro_attr_tk = new name{"[intro]"}; g_intro_attr_tk = new name{"[intro]"};
g_intro_bang_attr_tk = new name{"[intro!]"};
g_elim_attr_tk = new name{"[elim]"}; g_elim_attr_tk = new name{"[elim]"};
g_recursor_tk = new name{"[recursor"}; g_recursor_tk = new name{"[recursor"};
g_attribute_tk = new name{"attribute"}; g_attribute_tk = new name{"attribute"};
@ -442,10 +442,10 @@ void finalize_tokens() {
delete g_simp_attr_tk; delete g_simp_attr_tk;
delete g_congr_attr_tk; delete g_congr_attr_tk;
delete g_light_attr_tk; delete g_light_attr_tk;
delete g_backward_attr_tk;
delete g_no_pattern_attr_tk; delete g_no_pattern_attr_tk;
delete g_forward_attr_tk; delete g_forward_attr_tk;
delete g_intro_attr_tk; delete g_intro_attr_tk;
delete g_intro_bang_attr_tk;
delete g_elim_attr_tk; delete g_elim_attr_tk;
delete g_recursor_tk; delete g_recursor_tk;
delete g_attribute_tk; delete g_attribute_tk;
@ -602,10 +602,10 @@ name const & get_subst_tk() { return *g_subst_tk; }
name const & get_simp_attr_tk() { return *g_simp_attr_tk; } name const & get_simp_attr_tk() { return *g_simp_attr_tk; }
name const & get_congr_attr_tk() { return *g_congr_attr_tk; } name const & get_congr_attr_tk() { return *g_congr_attr_tk; }
name const & get_light_attr_tk() { return *g_light_attr_tk; } name const & get_light_attr_tk() { return *g_light_attr_tk; }
name const & get_backward_attr_tk() { return *g_backward_attr_tk; }
name const & get_no_pattern_attr_tk() { return *g_no_pattern_attr_tk; } name const & get_no_pattern_attr_tk() { return *g_no_pattern_attr_tk; }
name const & get_forward_attr_tk() { return *g_forward_attr_tk; } name const & get_forward_attr_tk() { return *g_forward_attr_tk; }
name const & get_intro_attr_tk() { return *g_intro_attr_tk; } name const & get_intro_attr_tk() { return *g_intro_attr_tk; }
name const & get_intro_bang_attr_tk() { return *g_intro_bang_attr_tk; }
name const & get_elim_attr_tk() { return *g_elim_attr_tk; } name const & get_elim_attr_tk() { return *g_elim_attr_tk; }
name const & get_recursor_tk() { return *g_recursor_tk; } name const & get_recursor_tk() { return *g_recursor_tk; }
name const & get_attribute_tk() { return *g_attribute_tk; } name const & get_attribute_tk() { return *g_attribute_tk; }

View file

@ -123,10 +123,10 @@ name const & get_subst_tk();
name const & get_simp_attr_tk(); name const & get_simp_attr_tk();
name const & get_congr_attr_tk(); name const & get_congr_attr_tk();
name const & get_light_attr_tk(); name const & get_light_attr_tk();
name const & get_backward_attr_tk();
name const & get_no_pattern_attr_tk(); name const & get_no_pattern_attr_tk();
name const & get_forward_attr_tk(); name const & get_forward_attr_tk();
name const & get_intro_attr_tk(); name const & get_intro_attr_tk();
name const & get_intro_bang_attr_tk();
name const & get_elim_attr_tk(); name const & get_elim_attr_tk();
name const & get_recursor_tk(); name const & get_recursor_tk();
name const & get_attribute_tk(); name const & get_attribute_tk();

View file

@ -116,10 +116,10 @@ subst [subst]
simp_attr [simp] simp_attr [simp]
congr_attr [congr] congr_attr [congr]
light_attr [light light_attr [light
backward_attr [backward]
no_pattern_attr [no_pattern] no_pattern_attr [no_pattern]
forward_attr [forward] forward_attr [forward]
intro_attr [intro] intro_attr [intro]
intro_bang_attr [intro!]
elim_attr [elim] elim_attr [elim]
recursor [recursor recursor [recursor
attribute attribute attribute attribute

View file

@ -1,11 +1,11 @@
constants (A B C : Prop) (H : A → B) (G : A → B → C) constants (A B C : Prop) (H : A → B) (G : A → B → C)
constants (T : Type) (f : T → A) constants (T : Type) (f : T → A)
attribute H [backward] attribute H [intro]
attribute G [backward] attribute G [intro]
attribute f [backward] attribute f [intro]
print H print H
print G print G
print f print f
print [backward] print [intro]

View file

@ -2,15 +2,15 @@
set_option blast.strategy "backward" set_option blast.strategy "backward"
constants {P Q R S T U : Prop} (Hpq : P → Q) (Hqr : Q → R) (Hrs : R → S) (Hst : S → T) constants {P Q R S T U : Prop} (Hpq : P → Q) (Hqr : Q → R) (Hrs : R → S) (Hst : S → T)
constants (Huq : U → Q) (Hur : U → R) (Hus : U → S) (Hut : U → T) constants (Huq : U → Q) (Hur : U → R) (Hus : U → S) (Hut : U → T)
attribute Hpq [backward] attribute Hpq [intro]
attribute Hqr [backward] attribute Hqr [intro]
attribute Hrs [backward] attribute Hrs [intro]
attribute Hst [backward] attribute Hst [intro]
attribute Huq [backward] attribute Huq [intro]
attribute Hur [backward] attribute Hur [intro]
attribute Hus [backward] attribute Hus [intro]
attribute Hut [backward] attribute Hut [intro]
definition lemma1 (p : P) : Q := by blast definition lemma1 (p : P) : Q := by blast
definition lemma2 (p : P) : R := by blast definition lemma2 (p : P) : R := by blast

View file

@ -2,10 +2,10 @@
set_option blast.strategy "backward" set_option blast.strategy "backward"
constants {P Q R S T U : Prop} constants {P Q R S T U : Prop}
constants (Huq : U → Q) (Hur : U → R) (Hus : U → S) (Hut : U → T) constants (Huq : U → Q) (Hur : U → R) (Hus : U → S) (Hut : U → T)
attribute Huq [backward] attribute Huq [intro]
attribute Hur [backward] attribute Hur [intro]
attribute Hus [backward] attribute Hus [intro]
attribute Hut [backward] attribute Hut [intro]
definition lemma1 : (P → Q) → P → Q := by blast definition lemma1 : (P → Q) → P → Q := by blast
definition lemma2 : (P → Q) → (Q → R) → P → R := by blast definition lemma2 : (P → Q) → (Q → R) → P → R := by blast

View file

@ -12,7 +12,7 @@ by blast
example : ∀ (F F' : Prop), F ∧ F' → F := example : ∀ (F F' : Prop), F ∧ F' → F :=
by blast by blast
attribute and.intro [backward] attribute and.intro [intro]
example : ∀ (P Q R : nat → Prop) (F : Prop), (F ∧ (∀ n m, (Q m ∧ R n) → P n)) → example : ∀ (P Q R : nat → Prop) (F : Prop), (F ∧ (∀ n m, (Q m ∧ R n) → P n)) →
(F → R 2) → Q 1 → P 2 ∧ F := (F → R 2) → Q 1 → P 2 ∧ F :=

View file

@ -9,7 +9,7 @@ constant subtype_refl : ∀ T, subtype T T
constant subtype_trans : ∀ S T U, subtype S T → subtype T U → subtype S U constant subtype_trans : ∀ S T U, subtype S T → subtype T U → subtype S U
attribute subtype_refl subtype_trans [backward] attribute subtype_refl subtype_trans [intro]
lemma L1 : ∀ T1 T2 T3 T4 T5 T6, subtype T1 T2 → subtype T2 T3 → subtype T3 T4 → subtype T4 T5 → subtype T5 T6 → subtype T1 T6 := lemma L1 : ∀ T1 T2 T3 T4 T5 T6, subtype T1 T2 → subtype T2 T3 → subtype T3 T4 → subtype T4 T5 → subtype T5 T6 → subtype T1 T6 :=
by blast by blast

View file

@ -11,7 +11,7 @@ by blast
print ex1 print ex1
attribute Exists.intro [intro] attribute Exists.intro [intro!] -- grind and core_grind only process [intro!] declarations
example (p q : nat → Prop) : (∃ x, p x ∧ q x) → (∃ x, q x) ∧ (∃ x, p x) := example (p q : nat → Prop) : (∃ x, p x ∧ q x) → (∃ x, q x) ∧ (∃ x, p x) :=
by blast by blast
@ -27,7 +27,5 @@ by blast
example (a : nat) : a = 0 → (λ x, x + a) = (λ x, x + 0) := example (a : nat) : a = 0 → (λ x, x + a) = (λ x, x + 0) :=
by blast by blast
attribute Exists.intro [intro]
example (p q : nat → Prop) : (∃ x, p x ∧ q x) → (∃ x, q x) ∧ (∃ x, p x) := example (p q : nat → Prop) : (∃ x, p x ∧ q x) → (∃ x, q x) ∧ (∃ x, p x) :=
by blast by blast