]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/reducibility/tif.ma
- predefined_virtuals: some additions
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / tif.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basic_2/reducibility/trf.ma".
16
17 (* CONTEXT-FREE IRREDUCIBLE TERMS *******************************************)
18
19 definition tif: predicate term ≝ λT. 𝐑⦃T⦄ → ⊥.
20
21 interpretation "context-free irreducibility (term)"
22    'NotReducible T = (tif T).
23
24 (* Basic inversion lemmas ***************************************************)
25
26 lemma tif_inv_abbr: ∀V,T. 𝐈⦃ⓓV.T⦄ → ⊥.
27 /2 width=1/ qed-.
28
29 lemma tif_inv_abst: ∀V,T. 𝐈⦃ⓛV.T⦄ → 𝐈⦃V⦄ ∧ 𝐈⦃T⦄.
30 /4 width=1/ qed-.
31
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/
36 qed-.
37
38 lemma tif_inv_cast: ∀V,T. 𝐈⦃ⓝV.T⦄ → ⊥.
39 /2 width=1/ qed-.
40
41 (* Basic properties *********************************************************)
42
43 lemma tif_atom: ∀I. 𝐈⦃⓪{I}⦄.
44 /2 width=4/ qed.
45
46 lemma tif_abst: ∀V,T. 𝐈⦃V⦄ →  𝐈⦃T⦄ →  𝐈⦃ⓛV.T⦄.
47 #V #T #HV #HT #H
48 elim (trf_inv_abst … H) -H /2 width=1/
49 qed.
50
51 lemma tif_appl: ∀V,T. 𝐈⦃V⦄ →  𝐈⦃T⦄ →  𝐒⦃T⦄ → 𝐈⦃ⓐV.T⦄.
52 #V #T #HV #HT #S #H
53 elim (trf_inv_appl … H) -H /2 width=1/
54 qed.