[ simplify in H1;
destruct H1
| simplify in H2;
- apply (bool_elim ? (eq n t));
+ apply (bool_elim ? (eq n a));
intro;
[ apply (ex_intro ? ? []);
apply (ex_intro ? ? l1);
elim (H H1 H2);
elim H4;
rewrite > H5;
- apply (ex_intro ? ? (t::a));
- apply (ex_intro ? ? a1);
+ apply (ex_intro ? ? (a::a1));
+ apply (ex_intro ? ? a2);
simplify;
reflexivity
]
[ simplify in H;
destruct H
| simplify in H1;
- apply (bool_elim ? (same_atom f t));
+ apply (bool_elim ? (same_atom f a));
intros;
[ elim (same_atom_to_exists ? ? H2);
autobatch
simplify;
reflexivity
| intros;
- generalize in match (refl_eq ? (mem ? same_atom t l2));
- elim (mem ? same_atom t l2) in ⊢ (? ? ? %→?);
+ generalize in match (refl_eq ? (mem ? same_atom a l2));
+ elim (mem ? same_atom a l2) in ⊢ (? ? ? %→?);
[ left;
elim (mem_to_exists_l1_l2 ? ? ? ? same_atom_to_eq H1);
elim H2; clear H2;
elim (mem_same_atom_to_exists ? ? H1);
rewrite > H2 in H3;
- apply (ex_intro ? ? a2);
+ apply (ex_intro ? ? a3);
rewrite > H2;
apply (ex_intro ? ? []);
simplify;
| elim (H l2);
[ left;
decompose;
- apply (ex_intro ? ? a);
- apply (ex_intro ? ? (t::a1));
+ apply (ex_intro ? ? a1);
+ apply (ex_intro ? ? (a::a2));
simplify;
- apply (ex_intro ? ? a2);
apply (ex_intro ? ? a3);
+ apply (ex_intro ? ? a4);
autobatch
| right;
intro;
- apply (bool_elim ? (same_atom t (FAtom n1)));
+ apply (bool_elim ? (same_atom a (FAtom n1)));
[ intro;
rewrite > (eq_to_eq_mem ? ? transitiveb_same_atom ? ? ? H3) in H1;
rewrite > H1;
autobatch
| intro;
change in ⊢ (? ? (? % ?) ?) with
- (match same_atom (FAtom n1) t with
+ (match same_atom (FAtom n1) a with
[true ⇒ true
|false ⇒ mem ? same_atom (FAtom n1) l
]);
destruct H2
| simplify;
intro;
- elim t;
+ elim a;
[ right;
apply (ex_intro ? ? []);
simplify;
elim (not_eq_nil_append_cons ? ? ? ? H6)
| elim H4;
right;
- apply (ex_intro ? ? (FFalse::a));
+ apply (ex_intro ? ? (FFalse::a1));
simplify;
elim H5;
- apply (ex_intro ? ? a1);
+ apply (ex_intro ? ? a2);
autobatch
|3,4: autobatch
| assumption
elim (not_eq_nil_append_cons ? ? ? ? H5)
| right;
elim H4;
- apply (ex_intro ? ? (FAtom n::a));
+ apply (ex_intro ? ? (FAtom n::a1));
simplify;
elim H;
autobatch
]
]
| intro;
- elim t;
+ elim a;
[ elim H;
[ left;
elim H4;
- apply (ex_intro ? ? (FTrue::a));
+ apply (ex_intro ? ? (FTrue::a1));
simplify;
elim H5;
autobatch
| elim H;
[ left;
elim H4;
- apply (ex_intro ? ? (FAtom n::a));
+ apply (ex_intro ? ? (FAtom n::a1));
simplify;
elim H5;
autobatch
elim S 1; clear S;
simplify in ⊢ (?→%→?);
intros;
- elim (look_for_axiom t t1);
+ elim (look_for_axiom a b);
[ decompose;
rewrite > H2; clear H2;
rewrite > H4; clear H4;
- apply (ExchangeL ? a1 a2 (FAtom a));
- apply (ExchangeR ? a3 a4 (FAtom a));
+ apply (ExchangeL ? a2 a3 (FAtom a1));
+ apply (ExchangeR ? a4 a5 (FAtom a1));
apply Axiom
- | elim (sizel_0_no_axiom_is_tautology t t1 H H1 H2);
+ | elim (sizel_0_no_axiom_is_tautology a b H H1 H2);
[ decompose;
rewrite > H3;
- apply (ExchangeL ? a a1 FFalse);
+ apply (ExchangeL ? a1 a2 FFalse);
apply FalseL
| decompose;
rewrite > H3;
- apply (ExchangeR ? a a1 FTrue);
+ apply (ExchangeR ? a1 a2 FTrue);
apply TrueR
]
]