]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/syntax/ac.ma
partial commit in static_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / ac.ma
index 9968544cb9dfff4ac0dc483b2bbe50251e6a002f..2d7d320e886ff7858be14ea5674683a6fa4d4ef1 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/lib/arith.ma".
+include "ground/arith/nat_le.ma".
+include "static_2/notation/functions/one_0.ma".
+include "static_2/notation/functions/two_0.ma".
+include "static_2/notation/functions/omega_0.ma".
 
-(* APPLICABILITY CONDITION  *************************************************)
+(* APPLICABILITY CONDITION **************************************************)
 
 (* applicability condition specification *)
 record ac: Type[0] ≝ {
-(* degree of the sort *)
-   appl: predicate nat
+(* applicability domain *)
+   ad: predicate nat
 }.
 
 (* applicability condition postulates *)
 record ac_props (a): Prop ≝ {
-   ac_dec: ∀m. Decidable (∃∃n. m ≤ n & appl a n)
+   ac_dec: ∀m. Decidable (∃∃n. ad a n & m ≤ n)
 }.
 
 (* Notable specifications ***************************************************)
@@ -33,15 +36,34 @@ definition apply_top: predicate nat ≝ λn. ⊤.
 
 definition ac_top: ac ≝ mk_ac apply_top.
 
+interpretation "any number (applicability domain)"
+  'Omega = (ac_top).
+
 lemma ac_top_props: ac_props ac_top ≝ mk_ac_props ….
 /3 width=3 by or_introl, ex2_intro/
 qed.
 
 definition ac_eq (k): ac ≝ mk_ac (eq … k).
 
+interpretation "one (applicability domain)"
+  'Two = (ac_eq (nsucc nzero)).
+
+interpretation "zero (applicability domain)"
+  'One = (ac_eq nzero).
+
 lemma ac_eq_props (k): ac_props (ac_eq k) ≝ mk_ac_props ….
-#m elim (le_dec m k) #Hm
+#m elim (nle_dec m k) #Hm
+[ /3 width=3 by or_introl, ex2_intro/
+| @or_intror * #n #Hn #Hmn destruct /2 width=1 by/
+]
+qed.
+
+definition ac_le (k): ac ≝ mk_ac (λn. n ≤ k).
+
+lemma ac_le_props (k): ac_props (ac_le k) ≝ mk_ac_props ….
+#m elim (nle_dec m k) #Hm
 [ /3 width=3 by or_introl, ex2_intro/
-| @or_intror * #n #Hmn #H destruct /2 width=1 by/
+| @or_intror * #n #Hn #Hmn
+  /3 width=3 by nle_trans/
 ]
 qed.