Pros: projections from concrete records are now automatically reduced away.
Cons: the name of the constructor is no longer pair, but mk_Prod.
inductive unit : Type[0] ≝ it: unit.
(* Prod *)
-inductive Prod (A,B:Type[0]) : Type[0] ≝
-pair : A → B → Prod A B.
+record Prod (A,B:Type[0]) : Type[0] ≝ {
+ fst: A
+ ; snd: B
+ }.
-interpretation "Pair construction" 'pair x y = (pair ? ? x y).
+interpretation "Pair construction" 'pair x y = (mk_Prod ? ? x y).
interpretation "Product" 'product x y = (Prod x y).
-definition fst ≝ λA,B.λp:A × B.
- match p with [pair a b ⇒ a].
-
-definition snd ≝ λA,B.λp:A × B.
- match p with [pair a b ⇒ b].
-
interpretation "pair pi1" 'pi1 = (fst ? ?).
interpretation "pair pi2" 'pi2 = (snd ? ?).
interpretation "pair pi1" 'pi1a x = (fst ? ? x).
lemma not_epsilon_sem: ∀S:DeqSet.∀a:S.∀w: word S. ∀e:pre S.
iff ((a::w) ∈ e) ((a::w) ∈ \fst e).
-#S #a #w * #i #b >fst_eq cases b normalize
+#S #a #w * #i #b cases b normalize
[% /2/ * // #H destruct |% normalize /2/]
qed.
@le_to_not_lt @sublist_length // * #e11 #e21 #membp
cut ((|\fst e11| = |\fst e1|) ∧ (|\fst e21| = |\fst e2|))
[|* #H1 #H2 <H1 <H2 @space_enum_complete]
- cases (H … membp) * #w * >fst_eq >snd_eq #we1 #we2 #_
+ cases (H … membp) * #w * #we1 #we2 #_
<we1 <we2 % //
|#m #HI * [#visited #vinv #finv >bisim_end //]
#p #front_tl #visited #Hn * #u_visited #vinv * #u_frontier #finv
cut (∀w.sublist ? w (occ S e1 e2)→∀p.memb (space S) p visited_res = true →
memb (space S) 〈moves ? w (\fst p), moves ? w (\snd p)〉 visited_res = true)
[#w elim w [//]
- #a #w1 #Hind #Hsub * #e11 #e21 #visp >fst_eq >snd_eq >moves_cons >moves_cons
+ #a #w1 #Hind #Hsub * #e11 #e21 #visp >moves_cons >moves_cons
@(Hind ? 〈?,?〉) [#x #H4 @Hsub @memb_cons //]
@(H1 〈?,?〉) // @Hsub @memb_hd] #all_reach
@bisim_char @occ_enough
// qed.
definition pre_concat_r ≝ λS:DeqSet.λi:pitem S.λe:pre S.
- match e with [ pair i1 b ⇒ 〈i · i1, b〉].
+ match e with [ mk_Prod i1 b ⇒ 〈i · i1, b〉].
notation "i ◂ e" left associative with precedence 60 for @{'ltrif $i $e}.
interpretation "pre_concat_r" 'ltrif i e = (pre_concat_r ? i e).
definition lc ≝ λS:DeqSet.λbcast:∀S:DeqSet.pitem S → pre S.λe1:pre S.λi2:pitem S.
match e1 with
- [ pair i1 b1 ⇒ match b1 with
+ [ mk_Prod i1 b1 ⇒ match b1 with
[ true ⇒ (i1 ◂ (bcast ? i2))
| false ⇒ 〈i1 · i2,false〉
]
definition lift ≝ λS.λf:pitem S →pre S.λe:pre S.
match e with
- [ pair i b ⇒ 〈\fst (f i), \snd (f i) ∨ b〉].
+ [ mk_Prod i b ⇒ 〈\fst (f i), \snd (f i) ∨ b〉].
notation "a ▸ b" left associative with precedence 60 for @{'lc eclose $a $b}.
interpretation "lc" 'lc op a b = (lc ? op a b).
definition lk ≝ λS:DeqSet.λbcast:∀S:DeqSet.∀E:pitem S.pre S.λe:pre S.
match e with
- [ pair i1 b1 ⇒
+ [ mk_Prod i1 b1 ⇒
match b1 with
[true ⇒ 〈(\fst (bcast ? i1))^*, true〉
|false ⇒ 〈i1^*,false〉
lemma erase_odot:∀S.∀e1,e2:pre S.
|\fst (e1 ⊙ e2)| = |\fst e1| · (|\fst e2|).
-#S * #i1 * * #i2 #b2 // >odot_true_b >fst_eq >fst_eq >fst_eq //
+#S * #i1 * * #i2 #b2 // >odot_true_b //
qed.
lemma ostar_true: ∀S.∀i:pitem S.
(**************************************************************************)
include "arithmetics/nat.ma".
-include "basics/list.ma".
+include "basics/lists/list.ma".
interpretation "iff" 'iff a b = (iff a b).