mirror of
https://github.com/achlipala/frap.git
synced 2024-11-10 00:07:51 +00:00
FirstClassFunctions: big old scary proof of sublistSummingToK_ok
This commit is contained in:
parent
a30079d6b4
commit
5019b2561e
1 changed files with 169 additions and 6 deletions
|
@ -20,10 +20,10 @@ Fixpoint filter {A : Set} (f : A -> bool) (ls : list A) : list A :=
|
|||
| x :: ls' => if f x then x :: filter f ls' else filter f ls'
|
||||
end.
|
||||
|
||||
Fixpoint foldl {A B : Set} (f : A -> B -> B) (acc : B) (ls : list A) : B :=
|
||||
Fixpoint fold_left {A B : Set} (f : B -> A -> B) (ls : list A) (acc : B) : B :=
|
||||
match ls with
|
||||
| nil => acc
|
||||
| x :: ls' => foldl f (f x acc) ls'
|
||||
| x :: ls' => fold_left f ls' (f acc x)
|
||||
end.
|
||||
|
||||
Record programming_language := {
|
||||
|
@ -66,8 +66,171 @@ Definition languages := [pascal; c; gallina; haskell; ocaml].
|
|||
|
||||
Compute map Name languages.
|
||||
Compute map Name (filter PurelyFunctional languages).
|
||||
Compute foldl max 0 (map AppearedInYear languages).
|
||||
Compute foldl max 0 (map AppearedInYear (filter PurelyFunctional languages)).
|
||||
Compute fold_left max (map AppearedInYear languages) 0.
|
||||
Compute fold_left max (map AppearedInYear (filter PurelyFunctional languages)) 0.
|
||||
|
||||
(* To avoid confusing things, we'll revert to the standard library's (identical)
|
||||
* versions of these functions for the remainder. *)
|
||||
Reset map.
|
||||
|
||||
|
||||
(** * Motivating continuations with search problems *)
|
||||
|
||||
Fixpoint allSublists {A : Set} (ls : list A) : list (list A) :=
|
||||
match ls with
|
||||
| [] => [[]]
|
||||
| x :: ls' =>
|
||||
let lss := allSublists ls' in
|
||||
lss ++ map (fun ls'' => x :: ls'') lss
|
||||
end.
|
||||
|
||||
Definition sum ls := fold_left plus ls 0.
|
||||
|
||||
Fixpoint sublistSummingTo (ns : list nat) (target : nat) : option (list nat) :=
|
||||
match filter (fun ns' => if sum ns' ==n target then true else false) (allSublists ns) with
|
||||
| ns' :: _ => Some ns'
|
||||
| [] => None
|
||||
end.
|
||||
|
||||
Fixpoint countingDown (from : nat) :=
|
||||
match from with
|
||||
| O => []
|
||||
| S from' => from' :: countingDown from'
|
||||
end.
|
||||
|
||||
Time Compute sublistSummingTo (countingDown 20) 1.
|
||||
|
||||
Fixpoint allSublistsK {A B : Set} (ls : list A)
|
||||
(failed : unit -> B)
|
||||
(found : list A -> (unit -> B) -> B) : B :=
|
||||
match ls with
|
||||
| [] => found [] failed
|
||||
| x :: ls' =>
|
||||
allSublistsK ls'
|
||||
failed
|
||||
(fun sol failed' =>
|
||||
found sol (fun _ => found (x :: sol) failed'))
|
||||
end.
|
||||
|
||||
Definition sublistSummingToK (ns : list nat) (target : nat) : option (list nat) :=
|
||||
allSublistsK ns
|
||||
(fun _ => None)
|
||||
(fun sol failed =>
|
||||
if sum sol ==n target then Some sol else failed tt).
|
||||
|
||||
Time Compute sublistSummingToK (countingDown 20) 1.
|
||||
|
||||
Theorem allSublistsK_ok : forall {A B : Set} (ls : list A) (failed : unit -> B) found,
|
||||
(forall sol, (exists ans, (forall failed', found sol failed' = ans)
|
||||
/\ ans <> failed tt)
|
||||
\/ (forall failed', found sol failed' = failed' tt))
|
||||
-> (exists sol ans, In sol (allSublists ls)
|
||||
/\ (forall failed', found sol failed' = ans)
|
||||
/\ allSublistsK ls failed found = ans
|
||||
/\ ans <> failed tt)
|
||||
\/ ((forall sol, In sol (allSublists ls)
|
||||
-> forall failed', found sol failed' = failed' tt)
|
||||
/\ allSublistsK ls failed found = failed tt).
|
||||
Proof.
|
||||
induct ls; simplify.
|
||||
|
||||
specialize (H []).
|
||||
first_order.
|
||||
right.
|
||||
propositional.
|
||||
subst.
|
||||
trivial.
|
||||
trivial.
|
||||
|
||||
assert (let found := (fun (sol : list A) (failed' : unit -> B) =>
|
||||
found sol (fun _ : unit => found (a :: sol) failed')) in
|
||||
(exists (sol : list A) (ans : B),
|
||||
In sol (allSublists ls) /\
|
||||
(forall failed' : unit -> B, found sol failed' = ans) /\
|
||||
allSublistsK ls failed found = ans /\ ans <> failed tt) \/
|
||||
(forall sol : list A,
|
||||
In sol (allSublists ls) -> forall failed' : unit -> B, found sol failed' = failed' tt) /\
|
||||
allSublistsK ls failed found = failed tt).
|
||||
apply IHls.
|
||||
first_order.
|
||||
generalize (H sol).
|
||||
first_order.
|
||||
specialize (H (a :: sol)).
|
||||
first_order.
|
||||
left.
|
||||
exists x; propositional.
|
||||
rewrite H0.
|
||||
trivial.
|
||||
right.
|
||||
simplify.
|
||||
rewrite H0.
|
||||
trivial.
|
||||
|
||||
clear IHls.
|
||||
simplify.
|
||||
first_order.
|
||||
|
||||
generalize (H x); first_order.
|
||||
left; exists x, x1; propositional.
|
||||
apply in_or_app; propositional.
|
||||
specialize (H1 failed).
|
||||
specialize (H4 (fun _ => found (a :: x) failed)).
|
||||
equality.
|
||||
left; exists (a :: x), x0; propositional.
|
||||
apply in_or_app; right; apply in_map_iff.
|
||||
first_order.
|
||||
specialize (H1 failed').
|
||||
rewrite H4 in H1.
|
||||
trivial.
|
||||
|
||||
right; propositional.
|
||||
apply in_app_or in H2; propositional.
|
||||
|
||||
generalize (H sol); first_order.
|
||||
apply H0 with (failed' := failed') in H3.
|
||||
rewrite H2 in H3.
|
||||
equality.
|
||||
|
||||
apply in_map_iff in H3.
|
||||
first_order.
|
||||
subst.
|
||||
generalize (H x); first_order.
|
||||
apply H0 with (failed' := failed) in H3.
|
||||
equality.
|
||||
apply H0 with (failed' := failed') in H3.
|
||||
rewrite H2 in H3; trivial.
|
||||
Qed.
|
||||
|
||||
Theorem sublistSummingToK_ok : forall ns target,
|
||||
match sublistSummingToK ns target with
|
||||
| None => forall sol, In sol (allSublists ns) -> sum sol <> target
|
||||
| Some sol => In sol (allSublists ns) /\ sum sol = target
|
||||
end.
|
||||
Proof.
|
||||
simplify.
|
||||
unfold sublistSummingToK.
|
||||
pose proof (allSublistsK_ok ns (fun _ => None)
|
||||
(fun sol failed => if sum sol ==n target then Some sol else failed tt)).
|
||||
cases H.
|
||||
|
||||
simplify.
|
||||
cases (sum sol ==n target).
|
||||
left; exists (Some sol); equality.
|
||||
propositional.
|
||||
|
||||
first_order.
|
||||
specialize (H0 (fun _ => None)).
|
||||
cases (sum x ==n target); try equality.
|
||||
subst.
|
||||
rewrite H1.
|
||||
propositional.
|
||||
|
||||
first_order.
|
||||
rewrite H0.
|
||||
simplify.
|
||||
apply H with (failed' := fun _ => None) in H1.
|
||||
cases (sum sol ==n target); equality.
|
||||
Qed.
|
||||
|
||||
|
||||
(** * The classics in continuation-passing style *)
|
||||
|
@ -84,10 +247,10 @@ Fixpoint filterK {A R : Set} (f : A -> (bool -> R) -> R) (ls : list A) (k : list
|
|||
| x :: ls' => f x (fun b => filterK f ls' (fun ls'' => k (if b then x :: ls'' else ls'')))
|
||||
end.
|
||||
|
||||
Fixpoint foldlK {A B R : Set} (f : A -> B -> (B -> R) -> R) (acc : B) (ls : list A) (k : B -> R) : R :=
|
||||
Fixpoint fold_lefK {A B R : Set} (f : B -> A -> (B -> R) -> R) (ls : list A) (acc : B) (k : B -> R) : R :=
|
||||
match ls with
|
||||
| nil => k acc
|
||||
| x :: ls' => f x acc (fun x' => foldlK f x' ls' k)
|
||||
| x :: ls' => f acc x (fun x' => foldl_leftK f ls' x' k)
|
||||
end.
|
||||
|
||||
Definition NameK {R : Set} (l : programming_language) (k : string -> R) : R :=
|
||||
|
|
Loading…
Reference in a new issue