From 18166618ab9029ed213ff3de556b5a9cba8673a8 Mon Sep 17 00:00:00 2001 From: matitaweb <claudio.sacerdoticoen@unibo.it> Date: Tue, 6 Mar 2012 11:44:26 +0000 Subject: [PATCH] chapter 9 and 10 --- weblib/tutorial/chapter10.ma | 326 +++++++++++++++++++++++++++++++++++ weblib/tutorial/chapter9.ma | 324 ---------------------------------- 2 files changed, 326 insertions(+), 324 deletions(-) create mode 100644 weblib/tutorial/chapter10.ma diff --git a/weblib/tutorial/chapter10.ma b/weblib/tutorial/chapter10.ma new file mode 100644 index 000000000..f52e80570 --- /dev/null +++ b/weblib/tutorial/chapter10.ma @@ -0,0 +1,326 @@ +include "cahpter9.ma". + +(* bisimulation *) +definition cofinal â λS.λp:(pre S)Ã(pre S). + \snd (\fst p) = \snd (\snd p). + +theorem equiv_sem: âS:DeqSet.âe1,e2:pre S. + \sem{e1} =1 \sem{e2} â âw.cofinal ? â©moves ? w e1,moves ? w e2âª. +#S #e1 #e2 % +[#same_sem #w + cut (âb1,b2. iff (b1 = true) (b2 = true) â (b1 = b2)) + [* * // * #H1 #H2 [@sym_eq @H1 //| @H2 //]] + #Hcut @Hcut @iff_trans [|@decidable_sem] + @iff_trans [|@same_sem] @iff_sym @decidable_sem +|#H #w1 @iff_trans [||@decidable_sem] <H @iff_sym @decidable_sem] +qed. + +definition occ â λS.λe1,e2:pre S. + unique_append ? (occur S (|\fst e1|)) (occur S (|\fst e2|)). + +lemma occ_enough: âS.âe1,e2:pre S. +(âw.(sublist S w (occ S e1 e2))â cofinal ? â©moves ? w e1,moves ? w e2âª) + ââw.cofinal ? â©moves ? w e1,moves ? w e2âª. +#S #e1 #e2 #H #w +cases (decidable_sublist S w (occ S e1 e2)) [@H] -H #H + >to_pit [2: @(not_to_not ⦠H) #H1 #a #memba @sublist_unique_append_l1 @H1 //] + >to_pit [2: @(not_to_not ⦠H) #H1 #a #memba @sublist_unique_append_l2 @H1 //] + // +qed. + +lemma equiv_sem_occ: âS.âe1,e2:pre S. +(âw.(sublist S w (occ S e1 e2))â cofinal ? â©moves ? w e1,moves ? w e2âª) +â \sem{e1}=1\sem{e2}. +#S #e1 #e2 #H @(proj2 ⦠(equiv_sem â¦)) @occ_enough #w @H +qed. + +definition sons â λS:DeqSet.λl:list S.λp:(pre S)Ã(pre S). + map ?? (λa.â©move S a (\fst (\fst p)),move S a (\fst (\snd p))âª) l. + +lemma memb_sons: âS,l.âp,q:(pre S)Ã(pre S). memb ? p (sons ? l q) = true â + âa.(move ? a (\fst (\fst q)) = \fst p ⧠+ move ? a (\fst (\snd q)) = \snd p). +#S #l elim l [#p #q normalize in ⢠(%â?); #abs @False_ind /2/] +#a #tl #Hind #p #q #H cases (orb_true_l ⦠H) -H + [#H @(ex_intro ⦠a) >(\P H) /2/ |#H @Hind @H] +qed. + +definition is_bisim â λS:DeqSet.λl:list ?.λalpha:list S. + âp:(pre S)Ã(pre S). memb ? p l = true â cofinal ? p ⧠(sublist ? (sons ? alpha p) l). + +lemma bisim_to_sem: âS:DeqSet.âl:list ?.âe1,e2: pre S. + is_bisim S l (occ S e1 e2) â memb ? â©e1,e2⪠l = true â \sem{e1}=1\sem{e2}. +#S #l #e1 #e2 #Hbisim #Hmemb @equiv_sem_occ +#w #Hsub @(proj1 ⦠(Hbisim â©moves S w e1,moves S w e2⪠?)) +lapply Hsub @(list_elim_left ⦠w) [//] +#a #w1 #Hind #Hsub >moves_left >moves_left @(proj2 â¦(Hbisim â¦(Hind ?))) + [#x #Hx @Hsub @memb_append_l1 // + |cut (memb S a (occ S e1 e2) = true) [@Hsub @memb_append_l2 //] #occa + @(memb_map ⦠occa) + ] +qed. + +(* the algorithm *) +let rec bisim S l n (frontier,visited: list ?) on n â + match n with + [ O â â©false,visited⪠(* assert false *) + | S m â + match frontier with + [ nil â â©true,visited⪠+ | cons hd tl â + if beqb (\snd (\fst hd)) (\snd (\snd hd)) then + bisim S l m (unique_append ? (filter ? (λx.notb (memb ? x (hd::visited))) + (sons S l hd)) tl) (hd::visited) + else â©false,visited⪠+ ] + ]. + +lemma unfold_bisim: âS,l,n.âfrontier,visited: list ?. + bisim S l n frontier visited = + match n with + [ O â â©false,visited⪠(* assert false *) + | S m â + match frontier with + [ nil â â©true,visited⪠+ | cons hd tl â + if beqb (\snd (\fst hd)) (\snd (\snd hd)) then + bisim S l m (unique_append ? (filter ? (λx.notb(memb ? x (hd::visited))) + (sons S l hd)) tl) (hd::visited) + else â©false,visited⪠+ ] + ]. +#S #l #n cases n // qed. + +lemma bisim_never: âS,l.âfrontier,visited: list ?. + bisim S l O frontier visited = â©false,visitedâª. +#frontier #visited >unfold_bisim // +qed. + +lemma bisim_end: âSig,l,m.âvisited: list ?. + bisim Sig l (S m) [] visited = â©true,visitedâª. +#n #visisted >unfold_bisim // +qed. + +lemma bisim_step_true: âSig,l,m.âp.âfrontier,visited: list ?. +beqb (\snd (\fst p)) (\snd (\snd p)) = true â + bisim Sig l (S m) (p::frontier) visited = + bisim Sig l m (unique_append ? (filter ? (λx.notb(memb ? x (p::visited))) + (sons Sig l p)) frontier) (p::visited). +#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // +qed. + +lemma bisim_step_false: âSig,l,m.âp.âfrontier,visited: list ?. +beqb (\snd (\fst p)) (\snd (\snd p)) = false â + bisim Sig l (S m) (p::frontier) visited = â©false,visitedâª. +#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // +qed. + +lemma notb_eq_true_l: âb. notb b = true â b = false. +#b cases b normalize // +qed. + +let rec pitem_enum S (i:re S) on i â + match i with + [ z â [pz S] + | e â [pe S] + | s y â [ps S y; pp S y] + | o i1 i2 â compose ??? (po S) (pitem_enum S i1) (pitem_enum S i2) + | c i1 i2 â compose ??? (pc S) (pitem_enum S i1) (pitem_enum S i2) + | k i â map ?? (pk S) (pitem_enum S i) + ]. + +lemma pitem_enum_complete : âS.âi:pitem S. + memb (DeqItem S) i (pitem_enum S (|i|)) = true. +#S #i elim i + [1,2:// + |3,4:#c normalize >(\b (refl ⦠c)) // + |5,6:#i1 #i2 #Hind1 #Hind2 @(memb_compose (DeqItem S) (DeqItem S)) // + |#i #Hind @(memb_map (DeqItem S)) // + ] +qed. + +definition pre_enum â λS.λi:re S. + compose ??? (λi,b.â©i,bâª) (pitem_enum S i) [true;false]. + +lemma pre_enum_complete : âS.âe:pre S. + memb ? e (pre_enum S (|\fst e|)) = true. +#S * #i #b @(memb_compose (DeqItem S) DeqBool ? (λi,b.â©i,bâª)) +// cases b normalize // +qed. + +definition space_enum â λS.λi1,i2:re S. + compose ??? (λe1,e2.â©e1,e2âª) (pre_enum S i1) (pre_enum S i2). + +lemma space_enum_complete : âS.âe1,e2: pre S. + memb ? â©e1,e2⪠(space_enum S (|\fst e1|) (|\fst e2|)) = true. +#S #e1 #e2 @(memb_compose ⦠(λi,b.â©i,bâª)) +// qed. + +definition all_reachable â λS.λe1,e2:pre S.λl: list ?. +uniqueb ? l = true ⧠+ âp. memb ? p l = true â + âw.(moves S w e1 = \fst p) ⧠(moves S w e2 = \snd p). + +definition disjoint â λS:DeqSet.λl1,l2. + âp:S. memb S p l1 = true â memb S p l2 = false. + +lemma bisim_correct: âS.âe1,e2:pre S.\sem{e1}=1\sem{e2} â + âl,n.âfrontier,visited:list ((pre S)Ã(pre S)). + |space_enum S (|\fst e1|) (|\fst e2|)| < n + |visited|â + all_reachable S e1 e2 visited â + all_reachable S e1 e2 frontier â + disjoint ? frontier visited â + \fst (bisim S l n frontier visited) = true. +#Sig #e1 #e2 #same #l #n elim n + [#frontier #visited #abs * #unique #H @False_ind @(absurd ⦠abs) + @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 * #we1 #we2 <we1 <we2 % >same_kernel_moves // + |#m #HI * [#visited #vinv #finv >bisim_end //] + #p #front_tl #visited #Hn * #u_visited #r_visited * #u_frontier #r_frontier + #disjoint + cut (âw.(moves ? w e1 = \fst p) ⧠(moves ? w e2 = \snd p)) + [@(r_frontier ⦠(memb_hd ⦠))] #rp + cut (beqb (\snd (\fst p)) (\snd (\snd p)) = true) + [cases rp #w * #fstp #sndp <fstp <sndp @(\b ?) + @(proj1 ⦠(equiv_sem ⦠)) @same] #ptest + >(bisim_step_true ⦠ptest) @HI -HI + [<plus_n_Sm // + |% [whd in ⢠(??%?); >(disjoint ⦠(memb_hd â¦)) whd in ⢠(??%?); // + |#p1 #H (cases (orb_true_l ⦠H)) [#eqp >(\P eqp) // |@r_visited] + ] + |whd % [@unique_append_unique @(andb_true_r ⦠u_frontier)] + @unique_append_elim #q #H + [cases (memb_sons ⦠(memb_filter_memb ⦠H)) -H + #a * #m1 #m2 cases rp #w1 * #mw1 #mw2 @(ex_intro ⦠(w1@[a])) + >moves_left >moves_left >mw1 >mw2 >m1 >m2 % // + |@r_frontier @memb_cons // + ] + |@unique_append_elim #q #H + [@injective_notb @(filter_true ⦠H) + |cut ((q==p) = false) + [|#Hpq whd in ⢠(??%?); >Hpq @disjoint @memb_cons //] + cases (andb_true ⦠u_frontier) #notp #_ @(\bf ?) + @(not_to_not ⦠not_eq_true_false) #eqqp <notp <eqqp >H // + ] + ] + ] +qed. + +definition all_true â λS.λl.âp:(pre S) à (pre S). memb ? p l = true â + (beqb (\snd (\fst p)) (\snd (\snd p)) = true). + +definition sub_sons â λS,l,l1,l2.âx:(pre S) à (pre S). +memb ? x l1 = true â sublist ? (sons ? l x) l2. + +lemma bisim_complete: + âS,l,n.âfrontier,visited,visited_res:list ?. + all_true S visited â + sub_sons S l visited (frontier@visited) â + bisim S l n frontier visited = â©true,visited_res⪠â + is_bisim S visited_res l ⧠sublist ? (frontier@visited) visited_res. +#S #l #n elim n + [#fron #vis #vis_res #_ #_ >bisim_never #H destruct + |#m #Hind * + [(* case empty frontier *) + -Hind #vis #vis_res #allv #H normalize in ⢠(%â?); + #H1 destruct % #p + [#membp % [@(\P ?) @allv //| @H //]|#H1 @H1] + |#hd cases (true_or_false (beqb (\snd (\fst hd)) (\snd (\snd hd)))) + [|(* case head of the frontier is non ok (absurd) *) + #H #tl #vis #vis_res #allv >(bisim_step_false ⦠H) #_ #H1 destruct] + (* frontier = hd:: tl and hd is ok *) + #H #tl #visited #visited_res #allv >(bisim_step_true ⦠H) + (* new_visited = hd::visited are all ok *) + cut (all_true S (hd::visited)) + [#p #H1 cases (orb_true_l ⦠H1) [#eqp >(\P eqp) @H |@allv]] + (* we now exploit the induction hypothesis *) + #allh #subH #bisim cases (Hind ⦠allh ⦠bisim) -bisim -Hind + [#H1 #H2 % // #p #membp @H2 -H2 cases (memb_append ⦠membp) -membp #membp + [cases (orb_true_l ⦠membp) -membp #membp + [@memb_append_l2 >(\P membp) @memb_hd + |@memb_append_l1 @sublist_unique_append_l2 // + ] + |@memb_append_l2 @memb_cons // + ] + |(* the only thing left to prove is the sub_sons invariant *) + #x #membx cases (orb_true_l ⦠membx) + [(* case x = hd *) + #eqhdx <(\P eqhdx) #xa #membxa + (* xa is a son of x; we must distinguish the case xa + was already visited form the case xa is new *) + cases (true_or_false ⦠(memb ? xa (x::visited))) + [(* xa visited - trivial *) #membxa @memb_append_l2 // + |(* xa new *) #membxa @memb_append_l1 @sublist_unique_append_l1 @memb_filter_l + [>membxa //|//] + ] + |(* case x in visited *) + #H1 #xa #membxa cases (memb_append ⦠(subH x ⦠H1 ⦠membxa)) + [#H2 (cases (orb_true_l ⦠H2)) + [#H3 @memb_append_l2 <(\P H3) @memb_hd + |#H3 @memb_append_l1 @sublist_unique_append_l2 @H3 + ] + |#H2 @memb_append_l2 @memb_cons @H2 + ] + ] + ] + ] +qed. + +definition equiv â λSig.λre1,re2:re Sig. + let e1 â â¢(blank ? re1) in + let e2 â â¢(blank ? re2) in + let n â S (length ? (space_enum Sig (|\fst e1|) (|\fst e2|))) in + let sig â (occ Sig e1 e2) in + (bisim ? sig n [â©e1,e2âª] []). + +theorem euqiv_sem : âSig.âe1,e2:re Sig. + \fst (equiv ? e1 e2) = true â \sem{e1} =1 \sem{e2}. +#Sig #re1 #re2 % + [#H @eqP_trans [|@eqP_sym @re_embedding] @eqP_trans [||@re_embedding] + cut (equiv ? re1 re2 = â©true,\snd (equiv ? re1 re2)âª) + [<H //] #Hcut + cases (bisim_complete ⦠Hcut) + [2,3: #p whd in ⢠((??%?)â?); #abs @False_ind /2/] + #Hbisim #Hsub @(bisim_to_sem ⦠Hbisim) + @Hsub @memb_hd + |#H @(bisim_correct ? (â¢(blank ? re1)) (â¢(blank ? re2))) + [@eqP_trans [|@re_embedding] @eqP_trans [|@H] @eqP_sym @re_embedding + |// + |% // #p whd in ⢠((??%?)â?); #abs @False_ind /2/ + |% // #p #H >(memb_single ⦠H) @(ex_intro ⦠ϵ) /2/ + |#p #_ normalize // + ] + ] +qed. + +lemma eqbnat_true : ân,m. eqbnat n m = true â n = m. +#n #m % [@eqbnat_true_to_eq | @eq_to_eqbnat_true] +qed. + +definition DeqNat â mk_DeqSet nat eqbnat eqbnat_true. + +definition a â s DeqNat O. +definition b â s DeqNat (S O). +definition c â s DeqNat (S (S O)). + +definition exp1 â ((a·b)^*·a). +definition exp2 â a·(b·a)^*. +definition exp4 â (b·a)^*. + +definition exp6 â a·(a ·a ·b^* + b^* ). +definition exp7 â a · a^* · b^*. + +definition exp8 â a·a·a·a·a·a·a·a·(a^* ). +definition exp9 â (a·a·a + a·a·a·a·a)^*. + +example ex1 : \fst (equiv ? (exp8+exp9) exp9) = true. +normalize // qed. + + + + + + + diff --git a/weblib/tutorial/chapter9.ma b/weblib/tutorial/chapter9.ma index dc7474e21..c7585df80 100644 --- a/weblib/tutorial/chapter9.ma +++ b/weblib/tutorial/chapter9.ma @@ -159,327 +159,3 @@ lemma to_pit: âS,w,e. ¬ sublist S w (occur S (|\fst e|)) â ] qed. -(* bisimulation *) -definition cofinal â λS.λp:(pre S)Ã(pre S). - \snd (\fst p) = \snd (\snd p). - -theorem equiv_sem: âS:DeqSet.âe1,e2:pre S. - \sem{e1} =1 \sem{e2} â âw.cofinal ? â©moves ? w e1,moves ? w e2âª. -#S #e1 #e2 % -[#same_sem #w - cut (âb1,b2. iff (b1 = true) (b2 = true) â (b1 = b2)) - [* * // * #H1 #H2 [@sym_eq @H1 //| @H2 //]] - #Hcut @Hcut @iff_trans [|@decidable_sem] - @iff_trans [|@same_sem] @iff_sym @decidable_sem -|#H #w1 @iff_trans [||@decidable_sem] <H @iff_sym @decidable_sem] -qed. - -definition occ â λS.λe1,e2:pre S. - unique_append ? (occur S (|\fst e1|)) (occur S (|\fst e2|)). - -lemma occ_enough: âS.âe1,e2:pre S. -(âw.(sublist S w (occ S e1 e2))â cofinal ? â©moves ? w e1,moves ? w e2âª) - ââw.cofinal ? â©moves ? w e1,moves ? w e2âª. -#S #e1 #e2 #H #w -cases (decidable_sublist S w (occ S e1 e2)) [@H] -H #H - >to_pit [2: @(not_to_not ⦠H) #H1 #a #memba @sublist_unique_append_l1 @H1 //] - >to_pit [2: @(not_to_not ⦠H) #H1 #a #memba @sublist_unique_append_l2 @H1 //] - // -qed. - -lemma equiv_sem_occ: âS.âe1,e2:pre S. -(âw.(sublist S w (occ S e1 e2))â cofinal ? â©moves ? w e1,moves ? w e2âª) -â \sem{e1}=1\sem{e2}. -#S #e1 #e2 #H @(proj2 ⦠(equiv_sem â¦)) @occ_enough #w @H -qed. - -definition sons â λS:DeqSet.λl:list S.λp:(pre S)Ã(pre S). - map ?? (λa.â©move S a (\fst (\fst p)),move S a (\fst (\snd p))âª) l. - -lemma memb_sons: âS,l.âp,q:(pre S)Ã(pre S). memb ? p (sons ? l q) = true â - âa.(move ? a (\fst (\fst q)) = \fst p ⧠- move ? a (\fst (\snd q)) = \snd p). -#S #l elim l [#p #q normalize in ⢠(%â?); #abs @False_ind /2/] -#a #tl #Hind #p #q #H cases (orb_true_l ⦠H) -H - [#H @(ex_intro ⦠a) >(\P H) /2/ |#H @Hind @H] -qed. - -definition is_bisim â λS:DeqSet.λl:list ?.λalpha:list S. - âp:(pre S)Ã(pre S). memb ? p l = true â cofinal ? p ⧠(sublist ? (sons ? alpha p) l). - -lemma bisim_to_sem: âS:DeqSet.âl:list ?.âe1,e2: pre S. - is_bisim S l (occ S e1 e2) â memb ? â©e1,e2⪠l = true â \sem{e1}=1\sem{e2}. -#S #l #e1 #e2 #Hbisim #Hmemb @equiv_sem_occ -#w #Hsub @(proj1 ⦠(Hbisim â©moves S w e1,moves S w e2⪠?)) -lapply Hsub @(list_elim_left ⦠w) [//] -#a #w1 #Hind #Hsub >moves_left >moves_left @(proj2 â¦(Hbisim â¦(Hind ?))) - [#x #Hx @Hsub @memb_append_l1 // - |cut (memb S a (occ S e1 e2) = true) [@Hsub @memb_append_l2 //] #occa - @(memb_map ⦠occa) - ] -qed. - -(* the algorithm *) -let rec bisim S l n (frontier,visited: list ?) on n â - match n with - [ O â â©false,visited⪠(* assert false *) - | S m â - match frontier with - [ nil â â©true,visited⪠- | cons hd tl â - if beqb (\snd (\fst hd)) (\snd (\snd hd)) then - bisim S l m (unique_append ? (filter ? (λx.notb (memb ? x (hd::visited))) - (sons S l hd)) tl) (hd::visited) - else â©false,visited⪠- ] - ]. - -lemma unfold_bisim: âS,l,n.âfrontier,visited: list ?. - bisim S l n frontier visited = - match n with - [ O â â©false,visited⪠(* assert false *) - | S m â - match frontier with - [ nil â â©true,visited⪠- | cons hd tl â - if beqb (\snd (\fst hd)) (\snd (\snd hd)) then - bisim S l m (unique_append ? (filter ? (λx.notb(memb ? x (hd::visited))) - (sons S l hd)) tl) (hd::visited) - else â©false,visited⪠- ] - ]. -#S #l #n cases n // qed. - -lemma bisim_never: âS,l.âfrontier,visited: list ?. - bisim S l O frontier visited = â©false,visitedâª. -#frontier #visited >unfold_bisim // -qed. - -lemma bisim_end: âSig,l,m.âvisited: list ?. - bisim Sig l (S m) [] visited = â©true,visitedâª. -#n #visisted >unfold_bisim // -qed. - -lemma bisim_step_true: âSig,l,m.âp.âfrontier,visited: list ?. -beqb (\snd (\fst p)) (\snd (\snd p)) = true â - bisim Sig l (S m) (p::frontier) visited = - bisim Sig l m (unique_append ? (filter ? (λx.notb(memb ? x (p::visited))) - (sons Sig l p)) frontier) (p::visited). -#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // -qed. - -lemma bisim_step_false: âSig,l,m.âp.âfrontier,visited: list ?. -beqb (\snd (\fst p)) (\snd (\snd p)) = false â - bisim Sig l (S m) (p::frontier) visited = â©false,visitedâª. -#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // -qed. - -lemma notb_eq_true_l: âb. notb b = true â b = false. -#b cases b normalize // -qed. - -let rec pitem_enum S (i:re S) on i â - match i with - [ z â [pz S] - | e â [pe S] - | s y â [ps S y; pp S y] - | o i1 i2 â compose ??? (po S) (pitem_enum S i1) (pitem_enum S i2) - | c i1 i2 â compose ??? (pc S) (pitem_enum S i1) (pitem_enum S i2) - | k i â map ?? (pk S) (pitem_enum S i) - ]. - -lemma pitem_enum_complete : âS.âi:pitem S. - memb (DeqItem S) i (pitem_enum S (|i|)) = true. -#S #i elim i - [1,2:// - |3,4:#c normalize >(\b (refl ⦠c)) // - |5,6:#i1 #i2 #Hind1 #Hind2 @(memb_compose (DeqItem S) (DeqItem S)) // - |#i #Hind @(memb_map (DeqItem S)) // - ] -qed. - -definition pre_enum â λS.λi:re S. - compose ??? (λi,b.â©i,bâª) (pitem_enum S i) [true;false]. - -lemma pre_enum_complete : âS.âe:pre S. - memb ? e (pre_enum S (|\fst e|)) = true. -#S * #i #b @(memb_compose (DeqItem S) DeqBool ? (λi,b.â©i,bâª)) -// cases b normalize // -qed. - -definition space_enum â λS.λi1,i2:re S. - compose ??? (λe1,e2.â©e1,e2âª) (pre_enum S i1) (pre_enum S i2). - -lemma space_enum_complete : âS.âe1,e2: pre S. - memb ? â©e1,e2⪠(space_enum S (|\fst e1|) (|\fst e2|)) = true. -#S #e1 #e2 @(memb_compose ⦠(λi,b.â©i,bâª)) -// qed. - -definition all_reachable â λS.λe1,e2:pre S.λl: list ?. -uniqueb ? l = true ⧠- âp. memb ? p l = true â - âw.(moves S w e1 = \fst p) ⧠(moves S w e2 = \snd p). - -definition disjoint â λS:DeqSet.λl1,l2. - âp:S. memb S p l1 = true â memb S p l2 = false. - -lemma bisim_correct: âS.âe1,e2:pre S.\sem{e1}=1\sem{e2} â - âl,n.âfrontier,visited:list ((pre S)Ã(pre S)). - |space_enum S (|\fst e1|) (|\fst e2|)| < n + |visited|â - all_reachable S e1 e2 visited â - all_reachable S e1 e2 frontier â - disjoint ? frontier visited â - \fst (bisim S l n frontier visited) = true. -#Sig #e1 #e2 #same #l #n elim n - [#frontier #visited #abs * #unique #H @False_ind @(absurd ⦠abs) - @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 * #we1 #we2 <we1 <we2 % >same_kernel_moves // - |#m #HI * [#visited #vinv #finv >bisim_end //] - #p #front_tl #visited #Hn * #u_visited #r_visited * #u_frontier #r_frontier - #disjoint - cut (âw.(moves ? w e1 = \fst p) ⧠(moves ? w e2 = \snd p)) - [@(r_frontier ⦠(memb_hd ⦠))] #rp - cut (beqb (\snd (\fst p)) (\snd (\snd p)) = true) - [cases rp #w * #fstp #sndp <fstp <sndp @(\b ?) - @(proj1 ⦠(equiv_sem ⦠)) @same] #ptest - >(bisim_step_true ⦠ptest) @HI -HI - [<plus_n_Sm // - |% [whd in ⢠(??%?); >(disjoint ⦠(memb_hd â¦)) whd in ⢠(??%?); // - |#p1 #H (cases (orb_true_l ⦠H)) [#eqp >(\P eqp) // |@r_visited] - ] - |whd % [@unique_append_unique @(andb_true_r ⦠u_frontier)] - @unique_append_elim #q #H - [cases (memb_sons ⦠(memb_filter_memb ⦠H)) -H - #a * #m1 #m2 cases rp #w1 * #mw1 #mw2 @(ex_intro ⦠(w1@[a])) - >moves_left >moves_left >mw1 >mw2 >m1 >m2 % // - |@r_frontier @memb_cons // - ] - |@unique_append_elim #q #H - [@injective_notb @(filter_true ⦠H) - |cut ((q==p) = false) - [|#Hpq whd in ⢠(??%?); >Hpq @disjoint @memb_cons //] - cases (andb_true ⦠u_frontier) #notp #_ @(\bf ?) - @(not_to_not ⦠not_eq_true_false) #eqqp <notp <eqqp >H // - ] - ] - ] -qed. - -definition all_true â λS.λl.âp:(pre S) à (pre S). memb ? p l = true â - (beqb (\snd (\fst p)) (\snd (\snd p)) = true). - -definition sub_sons â λS,l,l1,l2.âx:(pre S) à (pre S). -memb ? x l1 = true â sublist ? (sons ? l x) l2. - -lemma bisim_complete: - âS,l,n.âfrontier,visited,visited_res:list ?. - all_true S visited â - sub_sons S l visited (frontier@visited) â - bisim S l n frontier visited = â©true,visited_res⪠â - is_bisim S visited_res l ⧠sublist ? (frontier@visited) visited_res. -#S #l #n elim n - [#fron #vis #vis_res #_ #_ >bisim_never #H destruct - |#m #Hind * - [(* case empty frontier *) - -Hind #vis #vis_res #allv #H normalize in ⢠(%â?); - #H1 destruct % #p - [#membp % [@(\P ?) @allv //| @H //]|#H1 @H1] - |#hd cases (true_or_false (beqb (\snd (\fst hd)) (\snd (\snd hd)))) - [|(* case head of the frontier is non ok (absurd) *) - #H #tl #vis #vis_res #allv >(bisim_step_false ⦠H) #_ #H1 destruct] - (* frontier = hd:: tl and hd is ok *) - #H #tl #visited #visited_res #allv >(bisim_step_true ⦠H) - (* new_visited = hd::visited are all ok *) - cut (all_true S (hd::visited)) - [#p #H1 cases (orb_true_l ⦠H1) [#eqp >(\P eqp) @H |@allv]] - (* we now exploit the induction hypothesis *) - #allh #subH #bisim cases (Hind ⦠allh ⦠bisim) -bisim -Hind - [#H1 #H2 % // #p #membp @H2 -H2 cases (memb_append ⦠membp) -membp #membp - [cases (orb_true_l ⦠membp) -membp #membp - [@memb_append_l2 >(\P membp) @memb_hd - |@memb_append_l1 @sublist_unique_append_l2 // - ] - |@memb_append_l2 @memb_cons // - ] - |(* the only thing left to prove is the sub_sons invariant *) - #x #membx cases (orb_true_l ⦠membx) - [(* case x = hd *) - #eqhdx <(\P eqhdx) #xa #membxa - (* xa is a son of x; we must distinguish the case xa - was already visited form the case xa is new *) - cases (true_or_false ⦠(memb ? xa (x::visited))) - [(* xa visited - trivial *) #membxa @memb_append_l2 // - |(* xa new *) #membxa @memb_append_l1 @sublist_unique_append_l1 @memb_filter_l - [>membxa //|//] - ] - |(* case x in visited *) - #H1 #xa #membxa cases (memb_append ⦠(subH x ⦠H1 ⦠membxa)) - [#H2 (cases (orb_true_l ⦠H2)) - [#H3 @memb_append_l2 <(\P H3) @memb_hd - |#H3 @memb_append_l1 @sublist_unique_append_l2 @H3 - ] - |#H2 @memb_append_l2 @memb_cons @H2 - ] - ] - ] - ] -qed. - -definition equiv â λSig.λre1,re2:re Sig. - let e1 â â¢(blank ? re1) in - let e2 â â¢(blank ? re2) in - let n â S (length ? (space_enum Sig (|\fst e1|) (|\fst e2|))) in - let sig â (occ Sig e1 e2) in - (bisim ? sig n [â©e1,e2âª] []). - -theorem euqiv_sem : âSig.âe1,e2:re Sig. - \fst (equiv ? e1 e2) = true â \sem{e1} =1 \sem{e2}. -#Sig #re1 #re2 % - [#H @eqP_trans [|@eqP_sym @re_embedding] @eqP_trans [||@re_embedding] - cut (equiv ? re1 re2 = â©true,\snd (equiv ? re1 re2)âª) - [<H //] #Hcut - cases (bisim_complete ⦠Hcut) - [2,3: #p whd in ⢠((??%?)â?); #abs @False_ind /2/] - #Hbisim #Hsub @(bisim_to_sem ⦠Hbisim) - @Hsub @memb_hd - |#H @(bisim_correct ? (â¢(blank ? re1)) (â¢(blank ? re2))) - [@eqP_trans [|@re_embedding] @eqP_trans [|@H] @eqP_sym @re_embedding - |// - |% // #p whd in ⢠((??%?)â?); #abs @False_ind /2/ - |% // #p #H >(memb_single ⦠H) @(ex_intro ⦠ϵ) /2/ - |#p #_ normalize // - ] - ] -qed. - -lemma eqbnat_true : ân,m. eqbnat n m = true â n = m. -#n #m % [@eqbnat_true_to_eq | @eq_to_eqbnat_true] -qed. - -definition DeqNat â mk_DeqSet nat eqbnat eqbnat_true. - -definition a â s DeqNat O. -definition b â s DeqNat (S O). -definition c â s DeqNat (S (S O)). - -definition exp1 â ((a·b)^*·a). -definition exp2 â a·(b·a)^*. -definition exp4 â (b·a)^*. - -definition exp6 â a·(a ·a ·b^* + b^* ). -definition exp7 â a · a^* · b^*. - -definition exp8 â a·a·a·a·a·a·a·a·(a^* ). -definition exp9 â (a·a·a + a·a·a·a·a)^*. - -example ex1 : \fst (equiv ? (exp8+exp9) exp9) = true. -normalize // qed. - - - - - - - -- 2.39.5