]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs.ma
- the trace is explicit in all auto tactics with depth > 1
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / equivalence / cpcs.ma
index c979cedde9ec4ff3d86d04368f7c1483aa344882..cec1db1bc2e32a616a38a95014a95c8df0f28dbb 100644 (file)
@@ -45,7 +45,8 @@ lemma cpcs_refl: ∀G,L. reflexive … (cpcs G L).
 
 (* Basic_1: was: pc3_s *)
 lemma cpcs_sym: ∀G,L. symmetric … (cpcs G L).
-#G #L @TC_symmetric // qed-. (**) (* auto fails even after normalize *)
+normalize /3 width=1 by cpc_sym, TC_symmetric/
+qed-.
 
 lemma cpc_cpcs: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌ T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2.
 /2 width=1 by inj/ qed.