1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/reducibility/trf.ma".
17 (* CONTEXT-FREE IRREDUCIBLE TERMS *******************************************)
19 definition tif: predicate term ≝ λT. 𝐑⦃T⦄ → ⊥.
21 interpretation "context-free irreducibility (term)"
22 'NotReducible T = (tif T).
24 (* Basic inversion lemmas ***************************************************)
26 lemma tif_inv_abbr: ∀V,T. 𝐈⦃ⓓV.T⦄ → ⊥.
29 lemma tif_inv_abst: ∀V,T. 𝐈⦃ⓛV.T⦄ → 𝐈⦃V⦄ ∧ 𝐈⦃T⦄.
32 lemma tif_inv_appl: ∀V,T. 𝐈⦃ⓐV.T⦄ → ∧∧ 𝐈⦃V⦄ & 𝐈⦃T⦄ & 𝐒⦃T⦄.
33 #V #T #HVT @and3_intro /3 width=1/
34 generalize in match HVT; -HVT elim T -T //
35 * // * #U #T #_ #_ #H elim (H ?) -H /2 width=1/
38 lemma tif_inv_cast: ∀V,T. 𝐈⦃ⓝV.T⦄ → ⊥.
41 (* Basic properties *********************************************************)
43 lemma tif_atom: ∀I. 𝐈⦃⓪{I}⦄.
46 lemma tif_abst: ∀V,T. 𝐈⦃V⦄ → 𝐈⦃T⦄ → 𝐈⦃ⓛV.T⦄.
48 elim (trf_inv_abst … H) -H /2 width=1/
51 lemma tif_appl: ∀V,T. 𝐈⦃V⦄ → 𝐈⦃T⦄ → 𝐒⦃T⦄ → 𝐈⦃ⓐV.T⦄.
53 elim (trf_inv_appl … H) -H /2 width=1/