528 SOME vs => vs |
528 SOME vs => vs |
529 | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args)) |
529 | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args)) |
530 val _ = tracing "Preprocessing specification..." |
530 val _ = tracing "Preprocessing specification..." |
531 val T = Sign.the_const_type (ProofContext.theory_of ctxt) name |
531 val T = Sign.the_const_type (ProofContext.theory_of ctxt) name |
532 val t = Const (name, T) |
532 val t = Const (name, T) |
533 val ctxt' = ProofContext.theory (Context.copy_thy) ctxt |
533 val thy' = |
534 val thy' = Predicate_Compile.preprocess preprocess_options t (ProofContext.theory_of ctxt') |
534 Theory.copy (ProofContext.theory_of ctxt) |
535 val ctxt'' = ProofContext.init_global thy' |
535 |> Predicate_Compile.preprocess preprocess_options t |
|
536 val ctxt' = ProofContext.init_global thy' |
536 val _ = tracing "Generating prolog program..." |
537 val _ = tracing "Generating prolog program..." |
537 val (p, constant_table) = generate options ctxt'' name |
538 val (p, constant_table) = generate options ctxt' name |
538 |> (if #ensure_groundness options then add_ground_predicates ctxt'' else I) |
539 |> (if #ensure_groundness options then add_ground_predicates ctxt' else I) |
539 val _ = tracing "Running prolog program..." |
540 val _ = tracing "Running prolog program..." |
540 val tss = run p (translate_const constant_table name) (map first_upper vnames) soln |
541 val tss = run p (translate_const constant_table name) (map first_upper vnames) soln |
541 val _ = tracing "Restoring terms..." |
542 val _ = tracing "Restoring terms..." |
542 val empty = Const("Orderings.bot_class.bot", fastype_of t_compr) |
543 val empty = Const("Orderings.bot_class.bot", fastype_of t_compr) |
543 fun mk_insert x S = |
544 fun mk_insert x S = |
551 in |
552 in |
552 if null frees then |
553 if null frees then |
553 mk_set_compr (t :: in_insert) ts xs |
554 mk_set_compr (t :: in_insert) ts xs |
554 else |
555 else |
555 let |
556 let |
556 val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt'' [t]) ("uu", fastype_of t) |
557 val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt' [t]) ("uu", fastype_of t) |
557 val set_compr = |
558 val set_compr = |
558 HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t)) |
559 HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t)) |
559 frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"}))) |
560 frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"}))) |
560 in |
561 in |
561 mk_set_compr [] ts |
562 mk_set_compr [] ts |
562 (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)) |
563 (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)) |
563 end |
564 end |
564 end |
565 end |
565 in |
566 in |
566 foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr [] |
567 foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr [] |
567 (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt'' constant_table) (ts ~~ Ts))) tss) []) |
568 (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt' constant_table) (ts ~~ Ts))) tss) []) |
568 end |
569 end |
569 |
570 |
570 fun values_cmd print_modes soln raw_t state = |
571 fun values_cmd print_modes soln raw_t state = |
571 let |
572 let |
572 val ctxt = Toplevel.context_of state |
573 val ctxt = Toplevel.context_of state |
603 |
604 |
604 fun strip_horn A = (strip_imp_prems A, strip_imp_concl A); |
605 fun strip_horn A = (strip_imp_prems A, strip_imp_concl A); |
605 |
606 |
606 fun quickcheck ctxt report t size = |
607 fun quickcheck ctxt report t size = |
607 let |
608 let |
608 val ctxt' = ProofContext.theory (Context.copy_thy) ctxt |
609 val thy = Theory.copy (ProofContext.theory_of ctxt) |
609 val thy = (ProofContext.theory_of ctxt') |
|
610 val (vs, t') = strip_abs t |
610 val (vs, t') = strip_abs t |
611 val vs' = Variable.variant_frees ctxt' [] vs |
611 val vs' = Variable.variant_frees ctxt [] vs |
612 val Ts = map snd vs' |
612 val Ts = map snd vs' |
613 val t'' = subst_bounds (map Free (rev vs'), t') |
613 val t'' = subst_bounds (map Free (rev vs'), t') |
614 val (prems, concl) = strip_horn t'' |
614 val (prems, concl) = strip_horn t'' |
615 val constname = "quickcheck" |
615 val constname = "quickcheck" |
616 val full_constname = Sign.full_bname thy constname |
616 val full_constname = Sign.full_bname thy constname |
622 HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs'))) |
622 HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs'))) |
623 val tac = fn _ => Skip_Proof.cheat_tac thy1 |
623 val tac = fn _ => Skip_Proof.cheat_tac thy1 |
624 val intro = Goal.prove (ProofContext.init_global thy1) (map fst vs') [] t tac |
624 val intro = Goal.prove (ProofContext.init_global thy1) (map fst vs') [] t tac |
625 val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 |
625 val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 |
626 val thy3 = Predicate_Compile.preprocess preprocess_options const thy2 |
626 val thy3 = Predicate_Compile.preprocess preprocess_options const thy2 |
627 val ctxt'' = ProofContext.init_global thy3 |
627 val ctxt' = ProofContext.init_global thy3 |
628 val _ = tracing "Generating prolog program..." |
628 val _ = tracing "Generating prolog program..." |
629 val (p, constant_table) = generate {ensure_groundness = true} ctxt'' full_constname |
629 val (p, constant_table) = generate {ensure_groundness = true} ctxt' full_constname |
630 |> add_ground_predicates ctxt'' |
630 |> add_ground_predicates ctxt' |
631 val _ = tracing "Running prolog program..." |
631 val _ = tracing "Running prolog program..." |
632 val [ts] = run p (translate_const constant_table full_constname) (map fst vs') |
632 val [ts] = run p (translate_const constant_table full_constname) (map fst vs') |
633 (SOME 1) |
633 (SOME 1) |
634 val _ = tracing "Restoring terms..." |
634 val _ = tracing "Restoring terms..." |
635 val res = SOME (map (restore_term ctxt'' constant_table) (ts ~~ Ts)) |
635 val res = SOME (map (restore_term ctxt' constant_table) (ts ~~ Ts)) |
636 val empty_report = ([], false) |
636 val empty_report = ([], false) |
637 in |
637 in |
638 (res, empty_report) |
638 (res, empty_report) |
639 end; |
639 end; |
640 |
640 |