Require Automaton.
Require Import Alphabet.
Require Import List.
Require Import Syntax.
Module Make(
Import A:
Automaton.T).
We instantiate some sets/map. *
Module TerminalComparableM <:
ComparableM.
Definition t :=
terminal.
Instance tComparable :
Comparable t :=
_.
End TerminalComparableM.
Module TerminalOrderedType :=
OrderedType_from_ComparableM TerminalComparableM.
Module StateProdPosComparableM <:
ComparableM.
Definition t := (
state*
production*
nat)%
type.
Instance tComparable :
Comparable t :=
_.
End StateProdPosComparableM.
Module StateProdPosOrderedType :=
OrderedType_from_ComparableM StateProdPosComparableM.
Module TerminalSet :=
FSetAVL.Make TerminalOrderedType.
Module StateProdPosMap :=
FMapAVL.Make StateProdPosOrderedType.
Nullable predicate for symbols and list of symbols. *
Definition nullable_symb (
symbol:
symbol) :=
match symbol with
|
NT nt =>
nullable_nterm nt
|
_ =>
false
end.
Definition nullable_word (
word:
list symbol) :=
forallb nullable_symb word.
First predicate for non terminal, symbols and list of symbols, given as FSets. *
Definition first_nterm_set (
nterm:
nonterminal) :=
fold_left (
fun acc t =>
TerminalSet.add t acc)
(
first_nterm nterm)
TerminalSet.empty.
Definition first_symb_set (
symbol:
symbol) :=
match symbol with
|
NT nt =>
first_nterm_set nt
|
T t =>
TerminalSet.singleton t
end.
Fixpoint first_word_set (
word:
list symbol) :=
match word with
| [] =>
TerminalSet.empty
|
t::
q =>
if nullable_symb t then
TerminalSet.union (
first_symb_set t) (
first_word_set q)
else
first_symb_set t
end.
Small helper for finding the part of an item that is after the dot. *
Definition future_of_prod prod dot_pos :
list symbol :=
(
fix loop n lst :=
match n with
|
O =>
lst
|
S x =>
match loop x lst with [] => [] |
_::
q =>
q end
end)
dot_pos (
rev' (
prod_rhs_rev prod)).
We build a fast map to store all the items of all the states. *
Definition items_map (
_:
unit):
StateProdPosMap.t TerminalSet.t :=
fold_left (
fun acc state =>
fold_left (
fun acc item =>
let key := (
state,
prod_item item,
dot_pos_item item)
in
let data :=
fold_left (
fun acc t =>
TerminalSet.add t acc)
(
lookaheads_item item)
TerminalSet.empty
in
let old :=
match StateProdPosMap.find key acc with
|
Some x =>
x |
None =>
TerminalSet.empty
end
in
StateProdPosMap.add key (
TerminalSet.union data old)
acc
) (
items_of_state state)
acc
)
all_list (
StateProdPosMap.empty TerminalSet.t).
Accessor. *
Definition find_items_map items_map state prod dot_pos :
TerminalSet.t :=
match StateProdPosMap.find (
state,
prod,
dot_pos)
items_map with
|
None =>
TerminalSet.empty
|
Some x =>
x
end.
Definition state_has_future state prod (
fut:
list symbol) (
lookahead:
terminal) :=
exists dot_pos:
nat,
fut =
future_of_prod prod dot_pos /\
TerminalSet.In lookahead (
find_items_map (
items_map ())
state prod dot_pos).
Iterator over items. *
Definition forallb_items items_map (
P:
state ->
production ->
nat ->
TerminalSet.t ->
bool):
bool:=
StateProdPosMap.fold (
fun key set acc =>
match key with (
st,
p,
pos) => (
acc &&
P st p pos set)%
bool end
)
items_map true.
Lemma forallb_items_spec :
forall p,
forallb_items (
items_map ())
p =
true ->
forall st prod fut lookahead,
state_has_future st prod fut lookahead ->
forall P:
state ->
production ->
list symbol ->
terminal ->
Prop,
(
forall st prod pos set lookahead,
TerminalSet.In lookahead set ->
p st prod pos set =
true ->
P st prod (
future_of_prod prod pos)
lookahead) ->
P st prod fut lookahead.
Proof.
Validation for completeness *
The nullable predicate is a fixpoint : it is correct. *
Definition nullable_stable:=
forall p:
production,
nullable_word (
rev (
prod_rhs_rev p)) =
true ->
nullable_nterm (
prod_lhs p) =
true.
Definition is_nullable_stable (
_:
unit) :=
forallb (
fun p:
production =>
implb (
nullable_word (
rev' (
prod_rhs_rev p))) (
nullable_nterm (
prod_lhs p)))
all_list.
Property is_nullable_stable_correct :
is_nullable_stable () =
true ->
nullable_stable.
Proof.
The first predicate is a fixpoint : it is correct. *
Definition first_stable:=
forall (
p:
production),
TerminalSet.Subset (
first_word_set (
rev (
prod_rhs_rev p)))
(
first_nterm_set (
prod_lhs p)).
Definition is_first_stable (
_:
unit) :=
forallb (
fun p:
production =>
TerminalSet.subset (
first_word_set (
rev' (
prod_rhs_rev p)))
(
first_nterm_set (
prod_lhs p)))
all_list.
Property is_first_stable_correct :
is_first_stable () =
true ->
first_stable.
Proof.
The initial state has all the S=>.u items, where S is the start non-terminal *
Definition start_future :=
forall (
init:
initstate) (
t:
terminal) (
p:
production),
prod_lhs p =
start_nt init ->
state_has_future init p (
rev (
prod_rhs_rev p))
t.
Definition is_start_future items_map :=
forallb (
fun init =>
forallb (
fun prod =>
if compare_eqb (
prod_lhs prod) (
start_nt init)
then
let lookaheads :=
find_items_map items_map init prod 0
in
forallb (
fun t =>
TerminalSet.mem t lookaheads)
all_list
else
true)
all_list)
all_list.
Property is_start_future_correct :
is_start_future (
items_map ()) =
true ->
start_future.
Proof.
If a state contains an item of the form A->_.av[b], where a is a
terminal, then reading an a does a Shift_act, to a state containing
an item of the form A->_.v[b]. *
Definition terminal_shift :=
forall (
s1:
state)
prod fut lookahead,
state_has_future s1 prod fut lookahead ->
match fut with
|
T t::
q =>
match action_table s1 with
|
Lookahead_act awp =>
match awp t with
|
Shift_act s2 _ =>
state_has_future s2 prod q lookahead
|
_ =>
False
end
|
_ =>
False
end
|
_ =>
True
end.
Definition is_terminal_shift items_map :=
forallb_items items_map (
fun s1 prod pos lset =>
match future_of_prod prod pos with
|
T t::
_ =>
match action_table s1 with
|
Lookahead_act awp =>
match awp t with
|
Shift_act s2 _ =>
TerminalSet.subset lset (
find_items_map items_map s2 prod (
S pos))
|
_ =>
false
end
|
_ =>
false
end
|
_ =>
true
end).
Property is_terminal_shift_correct :
is_terminal_shift (
items_map ()) =
true ->
terminal_shift.
Proof.
If a state contains an item of the form A->_.[a], then either we do a
Default_reduce_act of the corresponding production, either a is a
terminal (ie. there is a lookahead terminal), and reading a does a
Reduce_act of the corresponding production. *
Definition end_reduce :=
forall (
s:
state)
prod fut lookahead,
state_has_future s prod fut lookahead ->
fut = [] ->
match action_table s with
|
Default_reduce_act p =>
p =
prod
|
Lookahead_act awt =>
match awt lookahead with
|
Reduce_act p =>
p =
prod
|
_ =>
False
end
end.
Definition is_end_reduce items_map :=
forallb_items items_map (
fun s prod pos lset =>
match future_of_prod prod pos with
| [] =>
match action_table s with
|
Default_reduce_act p =>
compare_eqb p prod
|
Lookahead_act awt =>
TerminalSet.fold (
fun lookahead acc =>
match awt lookahead with
|
Reduce_act p => (
acc &&
compare_eqb p prod)%
bool
|
_ =>
false
end)
lset true
end
|
_ =>
true
end).
Property is_end_reduce_correct :
is_end_reduce (
items_map ()) =
true ->
end_reduce.
Proof.
If a state contains an item of the form A->_.Bv[b], where B is a
non terminal, then the goto table says we have to go to a state containing
an item of the form A->_.v[b]. *
Definition non_terminal_goto :=
forall (
s1:
state)
prod fut lookahead,
state_has_future s1 prod fut lookahead ->
match fut with
|
NT nt::
q =>
match goto_table s1 nt with
|
Some (
exist _ s2 _) =>
state_has_future s2 prod q lookahead
|
None =>
forall prod fut lookahead,
state_has_future s1 prod fut lookahead ->
match fut with
|
NT nt'::
_ =>
nt <>
nt'
|
_ =>
True
end
end
|
_ =>
True
end.
Definition is_non_terminal_goto items_map :=
forallb_items items_map (
fun s1 prod pos lset =>
match future_of_prod prod pos with
|
NT nt::
_ =>
match goto_table s1 nt with
|
Some (
exist _ s2 _) =>
TerminalSet.subset lset (
find_items_map items_map s2 prod (
S pos))
|
None =>
forallb_items items_map (
fun s1'
prod'
pos'
_ =>
(
implb (
compare_eqb s1 s1')
match future_of_prod prod'
pos'
with
|
NT nt' ::
_ =>
negb (
compare_eqb nt nt')
|
_ =>
true
end)%
bool)
end
|
_ =>
true
end).
Property is_non_terminal_goto_correct :
is_non_terminal_goto (
items_map ()) =
true ->
non_terminal_goto.
Proof.
Definition start_goto :=
forall (
init:
initstate),
goto_table init (
start_nt init) =
None.
Definition is_start_goto (
_:
unit) :=
forallb (
fun (
init:
initstate) =>
match goto_table init (
start_nt init)
with
|
Some _ =>
false
|
None =>
true
end)
all_list.
Definition is_start_goto_correct:
is_start_goto () =
true ->
start_goto.
Proof.
Closure property of item sets : if a state contains an item of the form
A->_.Bv[b], then for each production B->u and each terminal a of
first(vb), the state contains an item of the form B->_.u[a] *
Definition non_terminal_closed :=
forall (
s1:
state)
prod fut lookahead,
state_has_future s1 prod fut lookahead ->
match fut with
|
NT nt::
q =>
forall (
p:
production) (
lookahead2:
terminal),
prod_lhs p =
nt ->
TerminalSet.In lookahead2 (
first_word_set q) \/
lookahead2 =
lookahead /\
nullable_word q =
true ->
state_has_future s1 p (
rev (
prod_rhs_rev p))
lookahead2
|
_ =>
True
end.
Definition is_non_terminal_closed items_map :=
forallb_items items_map (
fun s1 prod pos lset =>
match future_of_prod prod pos with
|
NT nt::
q =>
forallb (
fun p =>
if compare_eqb (
prod_lhs p)
nt then
let lookaheads :=
find_items_map items_map s1 p 0
in
(
implb (
nullable_word q) (
TerminalSet.subset lset lookaheads)) &&
TerminalSet.subset (
first_word_set q)
lookaheads
else true
)%
bool all_list
|
_ =>
true
end).
Property is_non_terminal_closed_correct:
is_non_terminal_closed (
items_map ()) =
true ->
non_terminal_closed.
Proof.
The automaton is complete *
Definition complete :=
nullable_stable /\
first_stable /\
start_future /\
terminal_shift
/\
end_reduce /\
non_terminal_goto /\
start_goto /\
non_terminal_closed.
Definition is_complete (
_:
unit) :=
let items_map :=
items_map ()
in
(
is_nullable_stable () &&
is_first_stable () &&
is_start_future items_map &&
is_terminal_shift items_map &&
is_end_reduce items_map &&
is_start_goto () &&
is_non_terminal_goto items_map &&
is_non_terminal_closed items_map)%
bool.
Property is_complete_correct:
is_complete () =
true ->
complete.
Proof.
End Make.