*)
-notation > "≔ (list0 ( ident x : T ) sep ,) opt (; (list1 (ident U ≟ term 90 V ) sep ,)) ⊢ term 19 Px ≡ term 19 Py"
+(* it seems unbelivable, but it works! *)
+notation > "≔ (list0 ( (list1 (ident x) sep , ) opt (: T) ) sep ,) opt (; (list1 (ident U ≟ term 90 V ) sep ,)) ⊢ term 19 Px ≡ term 19 Py"
with precedence 90
for @{ ${ fold right
@{ ${ default
@{ 'hint_decl $Px $Py }
}
}
- rec acc @{ ∀${ident x}:$T.$acc } } }.
+ rec acc @{
+ ${ fold right @{ $acc } rec acc2
+ @{ ∀${ident x}:${ default @{ $T } @{ ? } }.$acc2 } }
+ }
+ }}.
include "logic/pts.ma".
interpretation "hint_decl_CProp0" 'hint_decl a b = (hint_declaration_CProp0 ? a b).
interpretation "hint_decl_Type0" 'hint_decl a b = (hint_declaration_Type0 ? a b).
-(* little test
+(* little test
naxiom A : Type[0].
naxiom C : A → A → Type[0].
ndefinition D ≝ C.
alias symbol "hint_decl" = "hint_decl_Type1".
unification hint 0 ≔
- X : A, Y : A ; Z ≟ X, W ≟ Y ⊢ C X Y ≡ D Z W.
-
+ X, R : A, Y ; Z ≟ X, W ≟ Y ⊢ C X Y ≡ D Z W.
*)
| napply (H4 K2)]##]
nqed.
-unification hint 0 (∀A,B.(λx,y.True) (fun21 ??? and_morphism A B) (And A B)).
+unification hint 0 ≔ A,B ⊢ fun21 … and_morphism A B ≡ And A B.
(*nlemma test: ∀A,A',B: CProp[0]. A=A' → (B ∨ A) = B → (B ∧ A) ∧ B.
#A; #A'; #B; #H1; #H2;
| napply (H4 H)]##]
nqed.
-unification hint 0 (∀A,B.(λx,y.True) (fun21 ??? or_morphism A B) (Or A B)).
+unification hint 0 ≔ A,B ⊢ fun21 … or_morphism A B ≡ Or A B.
(*interpretation "or_morphism" 'or a b = (fun21 ??? or_morphism a b).*)
include "hints_declaration.ma".
alias symbol "hint_decl" = "hint_decl_Type2".
-unification hint 0 ≔ A : ? ⊢ carr1 (powerclass_setoid A) ≡ Ω^A.
+unification hint 0 ≔ A ⊢ carr1 (powerclass_setoid A) ≡ Ω^A.
(************ SETS OVER SETOIDS ********************)
| napply (qseteq A) ]
nqed.
-unification hint 0 ≔ A : ? ⊢
+unification hint 0 ≔ A ⊢
carr1 (qpowerclass_setoid A) ≡ qpowerclass A.
nlemma mem_ok: ∀A. binary_morphism1 (setoid1_of_setoid A) (qpowerclass_setoid A) CPROP.
nqed.
unification hint 0 ≔
- A : setoid, x : ?, S : ? ⊢ (mem_ok A) x S ≡ mem A S x.
+ A : setoid, x, S ⊢ (mem_ok A) x S ≡ mem A S x.
nlemma subseteq_ok: ∀A. binary_morphism1 (qpowerclass_setoid A) (qpowerclass_setoid A) CPROP.
#A; napply mk_binary_morphism1
(* unfold if intersect, exposing fun21 *)
alias symbol "hint_decl" = "hint_decl_Type1".
unification hint 0 ≔
- A : setoid, B : qpowerclass A, C : qpowerclass A ⊢
+ A : setoid, B,C : qpowerclass A ⊢
pc A (intersect_ok A B C) ≡ intersect ? (pc ? B) (pc ? C).
(* hints can pass under mem *) (* ??? XXX why is it needed? *)
-unification hint 0 ≔ A:?, B:?, x:?;
+unification hint 0 ≔ A,B,x ;
C ≟ B
(*---------------------*) ⊢
mem A B x ≡ mem A C x.