let E1' ≝ eclose ? E1 in
let E1'' ≝ snd … E1' in
match fst … E1' with
- [ true =>
+ [ true ⇒
let E2' ≝ eclose ? E2 in
〈 fst … E2', pc ? E1'' (snd … E2') 〉
| false ⇒ 〈 false, pc ? E1'' E2 〉 ]
| #w; nnormalize in ⊢ (??%%); napply (foo_nop false) ]##]
nqed.
+(*
nlet corec foo (a: unit):
equiv NAT
(eclose NAT (pc ? (ps ? 0) (pk ? (pc ? (ps ? 1) (ps ? 0)))))
##| #y; nnormalize in ⊢ (??%%); napply foo_nop
##]
nqed.
+*)
-ndefinition test1 ≝
- pc ? (pk ? (po ? (ps ? 0) (ps ? 1))) (ps ? 0).
-
-ndefinition test2 ≝
- po ?
- (pc ? (pk ? (pc ? (ps ? 0) (ps ? 1))) (ps ? 0))
- (pc ? (pk ? (pc ? (ps ? 0) (ps ? 1))) (ps ? 1)).
-
-ndefinition test3 ≝
- pk ? (pc ? (pc ? (ps ? 0) (pk ? (pc ? (ps ? 0) (ps ? 1)))) (ps ? 1)).
+notation < "a \sup ⋇" non associative with precedence 90 for @{ 'pk $a}.
+notation > "a ^ *" non associative with precedence 90 for @{ 'pk $a}.
+interpretation "star" 'pk a = (pk ? a).
+
+notation "❨a|b❩" non associative with precedence 90 for @{ 'po $a $b}.
+interpretation "or" 'po a b = (po ? a b).
+
+notation < "a b" non associative with precedence 60 for @{ 'pc $a $b}.
+notation > "a · b" non associative with precedence 60 for @{ 'pc $a $b}.
+interpretation "cat" 'pc a b = (pc ? a b).
+
+notation < "a" non associative with precedence 90 for @{ 'pp $a}.
+notation > "` term 90 a" non associative with precedence 90 for @{ 'pp $a}.
+interpretation "atom" 'pp a = (pp ? a).
+
+(* to get rid of \middot *)
+ncoercion rex_concat : ∀S:Type[0].∀p:pre S. pre S → pre S ≝ pc
+on _p : pre ? to ∀_:?.?.
+(* we could also get rid of ` with a coercion from nat → pre nat *)
+
+ndefinition test1 ≝ ❨ `0 | `1 ❩^* `0.
+ndefinition test2 ≝ ❨ (`0`1)^* `0 | (`0`1)^* `1 ❩.
+ndefinition test3 ≝ (`0 (`0`1)^* `1)^*.
nlemma foo: in_moves NAT
[0;0;1;0;1;1] (eclose ? test3) = true.
+ nnormalize in match test3;
nnormalize;
+//;
+nqed.
(**********************************************************)