definition cast ≝ λA:CProp.λa:A.a.
+notation < "\infrule (t\atop ⋮) a ?" with precedence 19 for @{ 'cast $a $t }.
+interpretation "cast" 'cast a t = (cast a t).
+
+definition assumpt ≝ λA:CProp.λa:A.a.
+
notation < "[ a ] \sup H" with precedence 19 for @{ 'ass $a $H }.
-interpretation "assumption" 'ass a H = (cast a H).
+interpretation "assumption" 'ass a H = (cast _ (assumpt a H)).
inductive Imply (A,B:CProp) : CProp ≝
Imply_intro: (A → B) → Imply A B.
axiom D: CProp.
axiom E: CProp.
+
+notation > "[H]" with precedence 90
+for @{ assumpt ? $H}.
+notation > "⇒\sub\i [ident H] term 90 b" with precedence 19
+for @{ Imply_intro ?? (λ${ident H}.cast $b ?) }.
+notation > "⇒\sub\e term 90 ab term 90 a" with precedence 19
+for @{ Imply_elim ?? (cast $ab ?) (cast $a ?) }.
+notation > "∧\sub\i term 90 a term 90 b" with precedence 19
+for @{ And_intro ?? (cast $a ?) (cast $b ?) }.
+notation > "∧\sub\e\sup\l term 90 ab" with precedence 19
+for @{ And_elim_l ?? (cast $ab ?) }.
+notation > "∧\sub\e\sup\r term 90 ab" with precedence 19
+for @{ And_elim_r ?? (cast $ab ?) }.
+notation > "∨\sub\i\sup\l term 90 a" with precedence 19
+for @{ Or_intro_l ?? (cast $a ?) }.
+notation > "∨\sub\i\sup\r term 90 a" with precedence 19
+for @{ Or_intro_r ?? (cast $a ?) }.
+notation > "∨\sub\e term 90 ab [ident Ha] term 90 c1 [ident Hb] term 90 c2" with precedence 19
+for @{ Or_elim ??? (cast $ab ?) (λ${ident Ha}.cast $c1 ?) (λ${ident Hb}.cast $c2 ?) }.
+
+
lemma ex1 : (A ⇒ E) ∨ B ⇒ A ∧ C ⇒ (E ∧ C) ∨ B.
-repeat (apply cast; constructor 1; intro);
-apply cast; apply (Or_elim (A ⇒ E) B (E∧C∨B)); try intro;
-[ apply cast; assumption
-| apply cast; apply Or_intro_l;
- apply cast; constructor 1;
- [ apply cast; apply (Imply_elim A E);
- [ apply cast; assumption
- | apply cast; apply (And_elim_l A C);
- apply cast; assumption
+ apply cast;
+ apply (⇒\sub\i [H] (A∧C⇒E∧C∨B));
+ apply (⇒\sub\i [K] (E∧C∨B));
+ apply (∨\sub\e ((A⇒E)∨B) [C1] (E∧C∨B) [C2] (E∧C∨B));
+[ apply [H];
+| apply (∨\sub\i\sup\l (E∧C));
+ apply (∧\sub\i E C);
+ [ apply (⇒\sub\e (A⇒E) A);
+ [ apply [C1];
+ | apply (∧\sub\e\sup\l (A∧C));
+ apply [K];
]
- | apply cast; apply (And_elim_r A C);
- apply cast; assumption
+ | apply (∧\sub\e\sup\r (A∧C));
+ apply [K];
]
-| apply cast; apply Or_intro_r;
- apply cast; assumption
+| apply (∨\sub\i\sup\r B);
+ apply [C2];
]
qed.