1 set "baseuri" "cic:/matita/tests/inversion/".
3 inductive nat : Set \def
7 inductive le (n:nat) : nat \to Prop \def
9 | leS : \forall m. le n m \to le n (S m).
11 alias symbol "eq" (instance 0) = "leibnitz's equality".
13 theorem test_inversion: \forall n. le n O \to n=O.
15 (* inversion begins *)
20 generalize in match Hcut.
21 apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
22 (* first goal (left open) *)
25 (* second goal (closed) *)
33 (* Piu' semplice e non lascia l'ipotesi inutile Hcut *)
34 alias id "refl_equal" = "cic:/Coq/Init/Logic/eq.ind#xpointer(1/1/1)".
35 theorem test_inversion2: \forall n. le n O \to n=O.
37 (* inversion begins *)
38 generalize in match (refl_equal nat O).
39 apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H).
40 (* first goal (left open) *)
42 (* second goal (closed) *)