src/HOL/ex/predicate_compile.ML
author bulwahn
Tue, 04 Aug 2009 08:34:56 +0200
changeset 32308 c2b74affab85
parent 32307 55166cd57a6d
child 32309 4a18f3cf6362
permissions -rw-r--r--
imported patch generic compilation of predicate compiler with different monads
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
     1
(* Author: Lukas Bulwahn, TU Muenchen
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     2
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     3
(Prototype of) A compiler from predicates specified by intro/elim rules
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     4
to equations.
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     5
*)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     6
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     7
signature PREDICATE_COMPILE =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     8
sig
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
     9
  type mode = int list option list * int list
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    10
  val add_equations_of: bool -> string list -> theory -> theory
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    11
  val register_predicate : (thm list * thm * int) -> theory -> theory
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
    12
  val is_registered : theory -> string -> bool
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
    13
  val fetch_pred_data : theory -> string -> (thm list * thm * int)  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    14
  val predfun_intro_of: theory -> string -> mode -> thm
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    15
  val predfun_elim_of: theory -> string -> mode -> thm
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    16
  val strip_intro_concl: int -> term -> term * (term list * term list)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    17
  val predfun_name_of: theory -> string -> mode -> string
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    18
  val all_preds_of : theory -> string list
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    19
  val modes_of: theory -> string -> mode list
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
    20
  val string_of_mode : mode -> string
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    21
  val intros_of: theory -> string -> thm list
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    22
  val nparams_of: theory -> string -> int
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
    23
  val add_intro: thm -> theory -> theory
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
    24
  val set_elim: thm -> theory -> theory
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
    25
  val setup: theory -> theory
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
    26
  val code_pred: string -> Proof.context -> Proof.state
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
    27
  val code_pred_cmd: string -> Proof.context -> Proof.state
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
    28
  val print_stored_rules: theory -> unit
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
    29
  val print_all_modes: theory -> unit
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    30
  val do_proofs: bool ref
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
    31
  val mk_casesrule : Proof.context -> int -> thm list -> term
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
    32
  val analyze_compr: theory -> term -> term
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
    33
  val eval_ref: (unit -> term Predicate.pred) option ref
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    34
  val add_equations : string -> theory -> theory
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
    35
  val code_pred_intros_attrib : attribute
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
    36
  (* used by Quickcheck_Generator *) 
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    37
  (*val funT_of : mode -> typ -> typ
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
    38
  val mk_if_pred : term -> term
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    39
  val mk_Eval : term * term -> term*)
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
    40
  val mk_tupleT : typ list -> typ
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    41
(*  val mk_predT :  typ -> typ *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    42
  (* temporary for compilation *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    43
  datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    44
  val prepare_intrs: theory -> string list ->
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    45
    (string * typ) list * int * string list * string list * (string * mode list) list *
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    46
    (string * (term list * indprem list) list) list * (string * (int option list * int)) list
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    47
  val infer_modes : theory -> (string * (int list option list * int list) list) list
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    48
    -> (string * (int option list * int)) list -> string list
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    49
    -> (string * (term list * indprem list) list) list -> (string * mode list) list
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
    50
  val split_mode : int list -> term list -> (term list * term list)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    51
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    52
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
    53
structure Predicate_Compile : PREDICATE_COMPILE =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    54
struct
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    55
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    56
(** auxiliary **)
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    57
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    58
(* debug stuff *)
31320
72eeb1b4e006 provide local dummy version of makestring -- NB: makestring is fragile and not portable, it should not occur in repository sources;
wenzelm
parents: 31284
diff changeset
    59
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    60
fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    61
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
    62
fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
    63
fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *)
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    64
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    65
val do_proofs = ref true;
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    66
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    67
fun mycheat_tac thy i st =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    68
  (Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) i) st
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    69
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
    70
fun remove_last_goal thy st =
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
    71
  (Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
    72
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    73
(* reference to preprocessing of InductiveSet package *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    74
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31580
diff changeset
    75
val ind_set_codegen_preproc = Inductive_Set.codegen_preproc;
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    76
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    77
(** fundamentals **)
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    78
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    79
(* syntactic operations *)
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    80
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    81
fun mk_eq (x, xs) =
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    82
  let fun mk_eqs _ [] = []
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    83
        | mk_eqs a (b::cs) =
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    84
            HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    85
  in mk_eqs x xs end;
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    86
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    87
fun mk_tupleT [] = HOLogic.unitT
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    88
  | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts;
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    89
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    90
fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = []
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    91
  | dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    92
  | dest_tupleT t = [t]
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
    93
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    94
fun mk_tuple [] = HOLogic.unit
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    95
  | mk_tuple ts = foldr1 HOLogic.mk_prod ts;
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    96
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    97
fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = []
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    98
  | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2)
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
    99
  | dest_tuple t = [t]
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   100
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   101
(** data structures for generic compilation for different monads **)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   102
(* maybe rename functions more generic:
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   103
  mk_predT -> mk_monadT; dest_predT -> dest_monadT
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   104
  mk_single -> mk_return (?)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   105
*)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   106
datatype compilation_funs = CompilationFuns of {
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   107
  mk_predT : typ -> typ,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   108
  dest_predT : typ -> typ,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   109
  mk_bot : typ -> term,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   110
  mk_single : term -> term,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   111
  mk_bind : term * term -> term,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   112
  mk_sup : term * term -> term,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   113
  mk_if : term -> term,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   114
  mk_not : term -> term
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   115
};
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   116
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   117
fun mk_predT (CompilationFuns funs) = #mk_predT funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   118
fun dest_predT (CompilationFuns funs) = #dest_predT funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   119
fun mk_bot (CompilationFuns funs) = #mk_bot funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   120
fun mk_single (CompilationFuns funs) = #mk_single funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   121
fun mk_bind (CompilationFuns funs) = #mk_bind funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   122
fun mk_sup (CompilationFuns funs) = #mk_sup funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   123
fun mk_if (CompilationFuns funs) = #mk_if funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   124
fun mk_not (CompilationFuns funs) = #mk_not funs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   125
  
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   126
structure PredicateCompFuns =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   127
struct
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   128
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   129
fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   130
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   131
fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   132
  | dest_predT T = raise TYPE ("dest_predT", [T], []);
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   133
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   134
fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   135
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   136
fun mk_single t =
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   137
  let val T = fastype_of t
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   138
  in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   139
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   140
fun mk_bind (x, f) =
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   141
  let val T as Type ("fun", [_, U]) = fastype_of f
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   142
  in
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   143
    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   144
  end;
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   145
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   146
val mk_sup = HOLogic.mk_binop @{const_name sup};
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   147
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   148
fun mk_if cond = Const (@{const_name Predicate.if_pred},
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   149
  HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   150
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   151
fun mk_not t = let val T = mk_predT HOLogic.unitT
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   152
  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   153
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   154
fun mk_Enum f =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   155
  let val T as Type ("fun", [T', _]) = fastype_of f
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   156
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   157
    Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   158
  end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   159
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   160
fun mk_Eval (f, x) =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   161
  let val T = fastype_of x
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   162
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   163
    Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   164
  end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   165
  
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   166
val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   167
  mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not} 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   168
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   169
end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   170
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   171
structure RPredCompFuns =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   172
struct
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   173
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   174
fun mk_rpredT T =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   175
  @{typ "Random.seed"} --> HOLogic.mk_prodT ((PredicateCompFuns.mk_predT T), @{typ "Random.seed"})
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   176
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   177
fun dest_rpredT (Type ("fun", [_,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   178
  Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   179
  | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   180
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   181
fun mk_rpredT T =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   182
  @{typ "Random.seed"} --> HOLogic.mk_prodT ((PredicateCompFuns.mk_predT T), @{typ "Random.seed"})
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   183
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   184
fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   185
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   186
fun mk_single t =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   187
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   188
    val T = fastype_of t
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   189
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   190
    Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   191
  end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   192
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   193
fun mk_bind (x, f) =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   194
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   195
    val T as (Type ("fun", [_, U])) = fastype_of f
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   196
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   197
    Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   198
  end
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   199
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   200
val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   201
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   202
fun mk_if cond = Const (@{const_name RPred.if_rpred},
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   203
  HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   204
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   205
fun mk_not t = error "Negation is not defined for RPred"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   206
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   207
val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   208
    mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not} 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   209
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   210
end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   211
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   212
(* destruction of intro rules *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   213
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   214
(* FIXME: look for other place where this functionality was used before *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   215
fun strip_intro_concl nparams intro = let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   216
  val _ $ u = Logic.strip_imp_concl intro
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   217
  val (pred, all_args) = strip_comb u
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   218
  val (params, args) = chop nparams all_args
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   219
in (pred, (params, args)) end
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   220
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   221
(* data structures *)
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   222
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   223
type mode = int list option list * int list; (*pmode FIMXE*)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   224
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   225
fun string_of_mode (iss, is) = space_implode " -> " (map
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   226
  (fn NONE => "X"
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   227
    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   228
       (iss @ [SOME is]));
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   229
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   230
fun print_modes modes = Output.tracing ("Inferred modes:\n" ^
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   231
  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   232
    string_of_mode ms)) modes));
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   233
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   234
datatype predfun_data = PredfunData of {
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   235
  name : string,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   236
  definition : thm,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   237
  intro : thm,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   238
  elim : thm
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   239
};
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   240
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   241
fun rep_predfun_data (PredfunData data) = data;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   242
fun mk_predfun_data (name, definition, intro, elim) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   243
  PredfunData {name = name, definition = definition, intro = intro, elim = elim}
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   244
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   245
datatype pred_data = PredData of {
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   246
  intros : thm list,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   247
  elim : thm option,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   248
  nparams : int,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   249
  functions : (mode * predfun_data) list
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   250
};
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   251
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   252
fun rep_pred_data (PredData data) = data;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   253
fun mk_pred_data ((intros, elim, nparams), functions) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   254
  PredData {intros = intros, elim = elim, nparams = nparams, functions = functions}
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   255
fun map_pred_data f (PredData {intros, elim, nparams, functions}) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   256
  mk_pred_data (f ((intros, elim, nparams), functions))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   257
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   258
fun eq_option eq (NONE, NONE) = true
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   259
  | eq_option eq (SOME x, SOME y) = eq (x, y)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   260
  | eq_option eq _ = false
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   261
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   262
fun eq_pred_data (PredData d1, PredData d2) = 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   263
  eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   264
  eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   265
  #nparams d1 = #nparams d2
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   266
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   267
structure PredData = TheoryDataFun
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   268
(
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   269
  type T = pred_data Graph.T;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   270
  val empty = Graph.empty;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   271
  val copy = I;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   272
  val extend = I;
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   273
  fun merge _ = Graph.merge eq_pred_data;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   274
);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   275
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   276
(* queries *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   277
31549
f7379ea2c949 tuned; corrected exception handling
bulwahn
parents: 31541
diff changeset
   278
fun lookup_pred_data thy name =
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   279
  Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   280
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   281
fun the_pred_data thy name = case lookup_pred_data thy name
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   282
 of NONE => error ("No such   predicate " ^ quote name)  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   283
  | SOME data => data;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   284
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   285
val is_registered = is_some oo lookup_pred_data 
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   286
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   287
val all_preds_of = Graph.keys o PredData.get
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   288
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   289
val intros_of = #intros oo the_pred_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   290
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   291
fun the_elim_of thy name = case #elim (the_pred_data thy name)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   292
 of NONE => error ("No elimination rule for predicate " ^ quote name)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   293
  | SOME thm => thm 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   294
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   295
val has_elim = is_some o #elim oo the_pred_data;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   296
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   297
val nparams_of = #nparams oo the_pred_data
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   298
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   299
val modes_of = (map fst) o #functions oo the_pred_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   300
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   301
fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   302
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   303
val is_compiled = not o null o #functions oo the_pred_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   304
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   305
fun lookup_predfun_data thy name mode =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   306
  Option.map rep_predfun_data (AList.lookup (op =)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   307
  (#functions (the_pred_data thy name)) mode)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   308
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   309
fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   310
 of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   311
  | SOME data => data;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   312
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   313
val predfun_name_of = #name ooo the_predfun_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   314
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   315
val predfun_definition_of = #definition ooo the_predfun_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   316
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   317
val predfun_intro_of = #intro ooo the_predfun_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   318
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   319
val predfun_elim_of = #elim ooo the_predfun_data
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   320
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   321
(* TODO: maybe join chop nparams and split_mode is
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   322
to some function split_mode mode and rename split_mode to split_smode *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   323
fun split_mode is ts = let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   324
  fun split_mode' _ _ [] = ([], [])
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   325
    | split_mode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   326
        (split_mode' is (i+1) ts)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   327
in split_mode' is 1 ts end
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   328
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   329
(* Remark: types of param_funT_of and funT_of are swapped - which is the more
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   330
canonical order? *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   331
(* maybe remove param_funT_of completely - by using funT_of *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   332
fun param_funT_of compfuns T NONE = T
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   333
  | param_funT_of compfuns T (SOME mode) = let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   334
     val Ts = binder_types T;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   335
     val (Us1, Us2) = split_mode mode Ts
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   336
   in Us1 ---> (mk_predT compfuns (mk_tupleT Us2)) end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   337
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   338
fun funT_of compfuns (iss, is) T = let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   339
    val Ts = binder_types T
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   340
    val (paramTs, argTs) = chop (length iss) Ts
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   341
    val paramTs' = map2 (fn SOME is => funT_of compfuns ([], is) | NONE => I) iss paramTs 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   342
    val (inargTs, outargTs) = split_mode is argTs
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   343
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   344
    (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   345
  end; 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   346
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   347
(* TODO: duplicate code in funT_of and this function *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   348
fun mk_predfun_of thy compfuns (name, T) mode = 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   349
  Const (predfun_name_of thy name mode, funT_of compfuns mode T)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   350
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   351
fun print_stored_rules thy =
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   352
  let
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   353
    val preds = (Graph.keys o PredData.get) thy
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   354
    fun print pred () = let
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   355
      val _ = writeln ("predicate: " ^ pred)
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   356
      val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   357
      val _ = writeln ("introrules: ")
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31986
diff changeset
   358
      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   359
        (rev (intros_of thy pred)) ()
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   360
    in
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   361
      if (has_elim thy pred) then
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31986
diff changeset
   362
        writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   363
      else
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   364
        writeln ("no elimrule defined")
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   365
    end
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   366
  in
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   367
    fold print preds ()
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   368
  end;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   369
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   370
fun print_all_modes thy =
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   371
  let
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   372
    val _ = writeln ("Inferred modes:")
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   373
    fun print (pred, modes) u =
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   374
      let
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   375
        val _ = writeln ("predicate: " ^ pred)
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   376
        val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   377
      in u end  
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   378
  in
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   379
    fold print (all_modes_of thy) ()
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   380
  end
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   381
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   382
(** preprocessing rules **)  
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   383
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   384
fun imp_prems_conv cv ct =
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   385
  case Thm.term_of ct of
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   386
    Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   387
  | _ => Conv.all_conv ct
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   388
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   389
fun Trueprop_conv cv ct =
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   390
  case Thm.term_of ct of
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   391
    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   392
  | _ => error "Trueprop_conv"
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   393
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   394
fun preprocess_intro thy rule =
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   395
  Conv.fconv_rule
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   396
    (imp_prems_conv
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   397
      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   398
    (Thm.transfer thy rule)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   399
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   400
fun preprocess_elim thy nargs elimrule = let
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   401
   fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   402
      HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   403
    | replace_eqs t = t
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   404
   fun preprocess_case t = let
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   405
     val params = Logic.strip_params t
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   406
     val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   407
     val assums_hyp' = assums1 @ (map replace_eqs assums2)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   408
     in list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t)) end
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   409
   val prems = Thm.prems_of elimrule
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   410
   val cases' = map preprocess_case (tl prems)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   411
   val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   412
 in
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   413
   Thm.equal_elim
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   414
     (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   415
        (cterm_of thy elimrule')))
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   416
     elimrule
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   417
 end;
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   418
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   419
(* special case: predicate with no introduction rule *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   420
fun noclause thy predname = let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   421
  val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   422
  val Ts = binder_types T
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   423
  val names = Name.variant_list []
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   424
        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   425
  val vs = map2 (curry Free) names Ts
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   426
  val clausehd =  HOLogic.mk_Trueprop (list_comb(Const (predname, T), vs))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   427
  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   428
  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   429
  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   430
  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   431
        (fn {...} => etac @{thm FalseE} 1)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   432
  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   433
        (fn {...} => etac (the_elim_of thy predname) 1) 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   434
in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   435
  ([intro], elim, 0)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   436
end
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   437
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   438
fun fetch_pred_data thy name =
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31580
diff changeset
   439
  case try (Inductive.the_inductive (ProofContext.init thy)) name of
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   440
    SOME (info as (_, result)) => 
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   441
      let
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   442
        fun is_intro_of intro =
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   443
          let
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   444
            val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   445
          in (fst (dest_Const const) = name) end;
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   446
        val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result)) 
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   447
        val elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31580
diff changeset
   448
        val nparams = length (Inductive.params_of (#raw_induct result))
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   449
      in if null intros then noclause thy name else (intros, elim, nparams) end                                                                    
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   450
  | NONE => error ("No such predicate: " ^ quote name)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   451
  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   452
(* updaters *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   453
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   454
fun add_predfun name mode data = let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   455
    val add = apsnd (cons (mode, mk_predfun_data data))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   456
  in PredData.map (Graph.map_node name (map_pred_data add)) end
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   457
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   458
fun is_inductive_predicate thy name =
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   459
  is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   460
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   461
fun depending_preds_of thy intros = fold Term.add_consts (map Thm.prop_of intros) [] |> map fst
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   462
    |> filter (fn c => is_inductive_predicate thy c orelse is_registered thy c)
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   463
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   464
(* code dependency graph *)    
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   465
fun dependencies_of thy name =
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   466
  let
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   467
    val (intros, elim, nparams) = fetch_pred_data thy name 
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   468
    val data = mk_pred_data ((intros, SOME elim, nparams), [])
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   469
    val keys = depending_preds_of thy intros
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   470
  in
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   471
    (data, keys)
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   472
  end;
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   473
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   474
(* TODO: add_edges - by analysing dependencies *)
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   475
fun add_intro thm thy = let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   476
   val (name, _) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   477
   fun cons_intro gr =
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   478
     case try (Graph.get_node gr) name of
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   479
       SOME pred_data => Graph.map_node name (map_pred_data
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   480
         (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   481
     | NONE =>
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   482
       let
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   483
         val nparams = the_default 0 (try (#3 o fetch_pred_data thy) name)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   484
       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), [])) gr end;
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   485
  in PredData.map cons_intro thy end
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   486
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   487
fun set_elim thm = let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   488
    val (name, _) = dest_Const (fst 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   489
      (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   490
    fun set (intros, _, nparams) = (intros, SOME thm, nparams)  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   491
  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   492
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   493
fun set_nparams name nparams = let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   494
    fun set (intros, elim, _ ) = (intros, elim, nparams) 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   495
  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   496
    
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   497
fun register_predicate (intros, elim, nparams) thy = let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   498
    val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd intros))))
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   499
  in
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   500
    PredData.map (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), []))
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   501
      #> fold Graph.add_edge (map (pair name) (depending_preds_of thy intros))) thy
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
   502
  end
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   503
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   504
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   505
(* Mode analysis *)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   506
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   507
(*** check if a term contains only constructor functions ***)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   508
fun is_constrt thy =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   509
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   510
    val cnstrs = flat (maps
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   511
      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
   512
      (Symtab.dest (Datatype.get_all thy)));
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   513
    fun check t = (case strip_comb t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   514
        (Free _, []) => true
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   515
      | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   516
            (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   517
          | _ => false)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   518
      | _ => false)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   519
  in check end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   520
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   521
(*** check if a type is an equality type (i.e. doesn't contain fun)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   522
  FIXME this is only an approximation ***)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   523
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   524
  | is_eqT _ = true;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   525
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   526
fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   527
val terms_vs = distinct (op =) o maps term_vs;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   528
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   529
(** collect all Frees in a term (with duplicates!) **)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   530
fun term_vTs tm =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   531
  fold_aterms (fn Free xT => cons xT | _ => I) tm [];
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   532
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   533
(*FIXME this function should not be named merge... make it local instead*)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   534
fun merge xs [] = xs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   535
  | merge [] ys = ys
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   536
  | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   537
      else y::merge (x::xs) ys;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   538
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   539
fun subsets i j = if i <= j then
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   540
       let val is = subsets (i+1) j
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   541
       in merge (map (fn ks => i::ks) is) is end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   542
     else [[]];
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   543
     
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   544
(* FIXME: should be in library - map_prod *)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   545
fun cprod ([], ys) = []
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   546
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   547
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   548
fun cprods xss = foldr (map op :: o cprod) [[]] xss;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   549
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   550
datatype hmode = Mode of mode * int list * hmode option list; (*FIXME don't understand
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   551
  why there is another mode type tmode !?*)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   552
31515
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
   553
  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   554
(*TODO: cleanup function and put together with modes_of_term *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   555
fun modes_of_param default modes t = let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   556
    val (vs, t') = strip_abs t
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   557
    val b = length vs
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   558
    fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   559
        let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   560
          val (args1, args2) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   561
            if length args < length iss then
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   562
              error ("Too few arguments for inductive predicate " ^ name)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   563
            else chop (length iss) args;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   564
          val k = length args2;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   565
          val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   566
            (1 upto b)  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   567
          val partial_mode = (1 upto k) \\ perm
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   568
        in
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   569
          if not (partial_mode subset is) then [] else
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   570
          let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   571
            val is' = 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   572
            (fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm [])
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   573
            |> fold (fn i => if i > k then cons (i - k + b) else I) is
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   574
              
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   575
           val res = map (fn x => Mode (m, is', x)) (cprods (map
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   576
            (fn (NONE, _) => [NONE]
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   577
              | (SOME js, arg) => map SOME (filter
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   578
                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   579
                    (iss ~~ args1)))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   580
          in res end
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   581
        end)) (AList.lookup op = modes name)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   582
  in case strip_comb t' of
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   583
    (Const (name, _), args) => the_default default (mk_modes name args)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   584
    | (Var ((name, _), _), args) => the (mk_modes name args)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   585
    | (Free (name, _), args) => the (mk_modes name args)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   586
    | _ => default end
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   587
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   588
and modes_of_term modes t =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   589
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   590
    val ks = 1 upto length (binder_types (fastype_of t));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   591
    val default = [Mode (([], ks), ks, [])];
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   592
    fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   593
        let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   594
          val (args1, args2) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   595
            if length args < length iss then
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   596
              error ("Too few arguments for inductive predicate " ^ name)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   597
            else chop (length iss) args;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   598
          val k = length args2;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   599
          val prfx = 1 upto k
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   600
        in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   601
          if not (is_prefix op = prfx is) then [] else
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   602
          let val is' = map (fn i => i - k) (List.drop (is, k))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   603
          in map (fn x => Mode (m, is', x)) (cprods (map
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   604
            (fn (NONE, _) => [NONE]
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   605
              | (SOME js, arg) => map SOME (filter
31170
c6efe82fc652 added predicate transformation function for code generation
bulwahn
parents: 31156 31169
diff changeset
   606
                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   607
                    (iss ~~ args1)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   608
          end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   609
        end)) (AList.lookup op = modes name)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   610
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   611
  in
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   612
    case strip_comb (Envir.eta_contract t) of
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   613
      (Const (name, _), args) => the_default default (mk_modes name args)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   614
    | (Var ((name, _), _), args) => the (mk_modes name args)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   615
    | (Free (name, _), args) => the (mk_modes name args)
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   616
    | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   617
    | _ => default
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   618
  end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   619
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   620
datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   621
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   622
fun print_clausess thy clausess =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   623
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   624
    val _ = Output.tracing "function print_clauses"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   625
    fun print_prem (Prem (ts, p)) = Syntax.string_of_term_global thy (list_comb (p, ts)) 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   626
      | print_prem _ = error "print_clausess: unimplemented"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   627
    fun print_clause pred (ts, prems) =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   628
      (space_implode " --> " (map print_prem prems)) ^ " --> " ^ pred ^ " "
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   629
      ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   630
    fun print_clauses (pred, clauses) =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   631
      "clauses of " ^ pred ^ ": " ^ cat_lines (map (print_clause pred) clauses)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   632
    val _ = Output.tracing (cat_lines (map print_clauses clausess))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   633
  in () end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   634
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   635
fun select_mode_prem thy modes vs ps =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   636
  find_first (is_some o snd) (ps ~~ map
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   637
    (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   638
          let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   639
            val (in_ts, out_ts) = split_mode is us;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   640
            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   641
            val vTs = maps term_vTs out_ts';
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   642
            val dupTs = map snd (duplicates (op =) vTs) @
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   643
              List.mapPartial (AList.lookup (op =) vTs) vs;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   644
          in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   645
            terms_vs (in_ts @ in_ts') subset vs andalso
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   646
            forall (is_eqT o fastype_of) in_ts' andalso
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   647
            term_vs t subset vs andalso
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   648
            forall is_eqT dupTs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   649
          end)
31170
c6efe82fc652 added predicate transformation function for code generation
bulwahn
parents: 31156 31169
diff changeset
   650
            (modes_of_term modes t handle Option =>
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   651
               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   652
      | Negprem (us, t) => find_first (fn Mode (_, is, _) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   653
            length us = length is andalso
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   654
            terms_vs us subset vs andalso
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   655
            term_vs t subset vs)
31170
c6efe82fc652 added predicate transformation function for code generation
bulwahn
parents: 31156 31169
diff changeset
   656
            (modes_of_term modes t handle Option =>
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   657
               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   658
      | Sidecond t => if term_vs t subset vs then SOME (Mode (([], []), [], []))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   659
          else NONE
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   660
      ) ps);
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   661
      
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   662
fun check_mode_clause thy param_vs modes (iss, is) (ts, ps) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   663
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   664
    val modes' = modes @ List.mapPartial
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   665
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   666
        (param_vs ~~ iss); 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   667
    fun check_mode_prems vs [] = SOME vs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   668
      | check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   669
          NONE => NONE
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   670
        | SOME (x, _) => check_mode_prems
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   671
            (case x of Prem (us, _) => vs union terms_vs us | _ => vs)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   672
            (filter_out (equal x) ps))
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   673
    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_mode is ts));
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   674
    val in_vs = terms_vs in_ts;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   675
    val concl_vs = terms_vs ts
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   676
    val _ = Output.tracing ("ts :" ^ (commas (map (Syntax.string_of_term_global thy) ts)))
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   677
    val _ = ()
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   678
    val ret = forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   679
    forall (is_eqT o fastype_of) in_ts' andalso
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   680
    (case check_mode_prems (param_vs union in_vs) ps of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   681
       NONE => false
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   682
     | SOME vs => concl_vs subset vs)
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   683
  in
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
   684
    ret
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   685
  end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   686
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   687
fun check_modes_pred thy param_vs preds modes (p, ms) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   688
  let val SOME rs = AList.lookup (op =) preds p
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   689
  in (p, List.filter (fn m => case find_index
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   690
    (not o check_mode_clause thy param_vs modes m) rs of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   691
      ~1 => true
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   692
    | i => (Output.tracing ("Clause " ^ string_of_int (i+1) ^ " of " ^
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   693
      p ^ " violates mode " ^ string_of_mode m); false)) ms)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   694
  end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   695
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   696
fun fixp f (x : (string * mode list) list) =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   697
  let val y = f x
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   698
  in if x = y then x else fixp f y end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   699
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   700
fun infer_modes thy extra_modes arities param_vs preds = fixp (fn modes =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   701
  map (check_modes_pred thy param_vs preds (modes @ extra_modes)) modes)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   702
    (map (fn (s, (ks, k)) => (s, cprod (cprods (map
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   703
      (fn NONE => [NONE]
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   704
        | SOME k' => map SOME (subsets 1 k')) ks),
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   705
      subsets 1 k))) arities);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   706
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   707
(* term construction *)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   708
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   709
fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   710
      NONE => (Free (s, T), (names, (s, [])::vs))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   711
    | SOME xs =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   712
        let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   713
          val s' = Name.variant names s;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   714
          val v = Free (s', T)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   715
        in
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   716
          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   717
        end);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   718
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   719
fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   720
  | distinct_v (t $ u) nvs =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   721
      let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   722
        val (t', nvs') = distinct_v t nvs;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   723
        val (u', nvs'') = distinct_v u nvs';
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   724
      in (t' $ u', nvs'') end
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   725
  | distinct_v x nvs = (x, nvs);
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   726
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   727
fun compile_match thy compfuns eqs eqs' out_ts success_t =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   728
  let 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   729
    val eqs'' = maps mk_eq eqs @ eqs'
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   730
    val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   731
    val name = Name.variant names "x";
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   732
    val name' = Name.variant (name :: names) "y";
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   733
    val T = mk_tupleT (map fastype_of out_ts);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   734
    val U = fastype_of success_t;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   735
    val U' = dest_predT compfuns U;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   736
    val v = Free (name, T);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   737
    val v' = Free (name', T);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   738
  in
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31580
diff changeset
   739
    lambda v (fst (Datatype.make_case
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   740
      (ProofContext.init thy) false [] v
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   741
      [(mk_tuple out_ts,
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   742
        if null eqs'' then success_t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   743
        else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   744
          foldr1 HOLogic.mk_conj eqs'' $ success_t $
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   745
            mk_bot compfuns U'),
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   746
       (v', mk_bot compfuns U')]))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   747
  end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   748
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   749
(*FIXME function can be removed*)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   750
fun mk_funcomp f t =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   751
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   752
    val names = Term.add_free_names t [];
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   753
    val Ts = binder_types (fastype_of t);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   754
    val vs = map Free
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   755
      (Name.variant_list names (replicate (length Ts) "x") ~~ Ts)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   756
  in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   757
    fold_rev lambda vs (f (list_comb (t, vs)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   758
  end;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   759
(*
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   760
fun compile_param_ext thy compfuns modes (NONE, t) = t
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   761
  | compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   762
      let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   763
        val (vs, u) = strip_abs t
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   764
        val (ivs, ovs) = split_mode is vs    
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   765
        val (f, args) = strip_comb u
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   766
        val (params, args') = chop (length ms) args
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   767
        val (inargs, outargs) = split_mode is' args'
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   768
        val b = length vs
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   769
        val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   770
        val outp_perm =
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   771
          snd (split_mode is perm)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   772
          |> map (fn i => i - length (filter (fn x => x < i) is'))
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   773
        val names = [] -- TODO
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   774
        val out_names = Name.variant_list names (replicate (length outargs) "x")
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   775
        val f' = case f of
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   776
            Const (name, T) =>
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   777
              if AList.defined op = modes name then
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   778
                mk_predfun_of thy compfuns (name, T) (iss, is')
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   779
              else error "compile param: Not an inductive predicate with correct mode"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   780
          | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is'))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   781
        val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f')))
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   782
        val out_vs = map Free (out_names ~~ outTs)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   783
        val params' = map (compile_param thy modes) (ms ~~ params)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   784
        val f_app = list_comb (f', params' @ inargs)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   785
        val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm)))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   786
        val match_t = compile_match thy compfuns [] [] out_vs single_t
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   787
      in list_abs (ivs,
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   788
        mk_bind compfuns (f_app, match_t))
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   789
      end
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   790
  | compile_param_ext _ _ _ _ = error "compile params"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   791
*)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   792
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   793
fun compile_param thy compfuns modes (NONE, t) = t
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   794
 | compile_param thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   795
   (* (case t of
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   796
     Abs _ => error "compile_param: Invalid term" *) (* compile_param_ext thy modes (m, t) *)
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   797
   (*  |  _ => let *)
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   798
   let  
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   799
     val (f, args) = strip_comb (Envir.eta_contract t)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   800
     val (params, args') = chop (length ms) args
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   801
     val params' = map (compile_param thy compfuns modes) (ms ~~ params)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   802
     val f' = case f of
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   803
        Const (name, T) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   804
          if AList.defined op = modes name then
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   805
             mk_predfun_of thy compfuns (name, T) (iss, is')
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   806
          else error "compile param: Not an inductive predicate with correct mode"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   807
      | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is'))
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
   808
   in list_comb (f', params' @ args') end
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   809
 | compile_param _ _ _ _ = error "compile params"
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   810
  
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   811
  
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   812
fun compile_expr thy compfuns modes (SOME (Mode (mode, is, ms)), t) =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   813
      (case strip_comb t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   814
         (Const (name, T), params) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   815
           if AList.defined op = modes name then
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   816
             let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   817
               val params' = map (compile_param thy compfuns modes) (ms ~~ params)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   818
             in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   819
               list_comb (mk_predfun_of thy compfuns (name, T) mode, params')
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   820
             end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   821
           else error "not a valid inductive expression"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   822
       | (Free (name, T), args) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   823
         (*if name mem param_vs then *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   824
         (* Higher order mode call *)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   825
         let val r = Free (name, param_funT_of compfuns T (SOME is))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   826
         in list_comb (r, args) end)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   827
  | compile_expr _ _ _ _ = error "not a valid inductive expression"
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   828
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   829
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   830
fun compile_clause thy compfuns all_vs param_vs modes (iss, is) (ts, ps) inp =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   831
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   832
    val modes' = modes @ List.mapPartial
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   833
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   834
        (param_vs ~~ iss);
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   835
    fun check_constrt t (names, eqs) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   836
      if is_constrt thy t then (t, (names, eqs)) else
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   837
        let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   838
          val s = Name.variant names "x";
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   839
          val v = Free (s, fastype_of t)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   840
        in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   841
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   842
    val (in_ts, out_ts) = split_mode is ts;
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   843
    val (in_ts', (all_vs', eqs)) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   844
      fold_map check_constrt in_ts (all_vs, []);
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   845
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   846
    fun compile_prems out_ts' vs names [] =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   847
          let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   848
            val (out_ts'', (names', eqs')) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   849
              fold_map check_constrt out_ts' (names, []);
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   850
            val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   851
              out_ts'' (names', map (rpair []) vs);
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   852
          in
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   853
            compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   854
              (mk_single compfuns (mk_tuple out_ts))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   855
          end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   856
      | compile_prems out_ts vs names ps =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   857
          let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   858
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   859
            val SOME (p, mode as SOME (Mode (_, js, _))) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   860
              select_mode_prem thy modes' vs' ps
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   861
            val ps' = filter_out (equal p) ps
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   862
            val (out_ts', (names', eqs)) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   863
              fold_map check_constrt out_ts (names, [])
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   864
            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   865
              out_ts' ((names', map (rpair []) vs))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   866
            val (compiled_clause, rest) = case p of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   867
               Prem (us, t) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   868
                 let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   869
                   val (in_ts, out_ts''') = split_mode js us;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   870
                   val u = list_comb (compile_expr thy compfuns modes (mode, t), in_ts)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   871
                   val rest = compile_prems out_ts''' vs' names'' ps'
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   872
                 in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   873
                   (u, rest)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   874
                 end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   875
             | Negprem (us, t) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   876
                 let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   877
                   val (in_ts, out_ts''') = split_mode js us
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   878
                   val u = list_comb (compile_expr thy compfuns modes (mode, t), in_ts)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   879
                   val rest = compile_prems out_ts''' vs' names'' ps'
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   880
                 in
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   881
                   (mk_not compfuns u, rest)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   882
                 end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   883
             | Sidecond t =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   884
                 let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   885
                   val rest = compile_prems [] vs' names'' ps';
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   886
                 in
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   887
                   (mk_if compfuns t, rest)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   888
                 end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   889
          in
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   890
            compile_match thy compfuns constr_vs' eqs out_ts'' 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   891
              (mk_bind compfuns (compiled_clause, rest))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   892
          end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   893
    val prem_t = compile_prems in_ts' param_vs all_vs' ps;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   894
  in
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   895
    mk_bind compfuns (mk_single compfuns inp, prem_t)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   896
  end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   897
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   898
fun compile_pred thy compfuns all_vs param_vs modes s T cls mode =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   899
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   900
    val Ts = binder_types T;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   901
    val (Ts1, Ts2) = chop (length param_vs) Ts;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   902
    val Ts1' = map2 (param_funT_of compfuns) Ts1 (fst mode)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   903
    val (Us1, _) = split_mode (snd mode) Ts2;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   904
    val xnames = Name.variant_list param_vs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   905
      (map (fn i => "x" ^ string_of_int i) (snd mode));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   906
    val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   907
    val cl_ts =
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   908
      map (fn cl => compile_clause thy compfuns
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   909
        all_vs param_vs modes mode cl (mk_tuple xs)) cls;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   910
  in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   911
    HOLogic.mk_Trueprop (HOLogic.mk_eq
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   912
      (list_comb (mk_predfun_of thy compfuns (s, T) mode,
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   913
         map2 (fn s => fn T => Free (s, T)) param_vs Ts1' @ xs),
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   914
       foldr1 (mk_sup compfuns) cl_ts))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   915
  end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   916
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   917
fun compile_preds thy compfuns all_vs param_vs modes preds =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   918
  map (fn (s, (T, cls)) =>
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   919
    map (compile_pred thy compfuns all_vs param_vs modes s T cls)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   920
      ((the o AList.lookup (op =) modes) s)) preds;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   921
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   922
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   923
(* special setup for simpset *)                  
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   924
val HOL_basic_ss' = HOL_basic_ss setSolver 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   925
  (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   926
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   927
(* Definition of executable functions and their intro and elim rules *)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   928
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30528
diff changeset
   929
fun print_arities arities = tracing ("Arities:\n" ^
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   930
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   931
    space_implode " -> " (map
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   932
      (fn NONE => "X" | SOME k' => string_of_int k')
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   933
        (ks @ [SOME k]))) arities));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   934
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   935
fun mk_Eval_of ((x, T), NONE) names = (x, names)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   936
  | mk_Eval_of ((x, T), SOME mode) names = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   937
  val Ts = binder_types T
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   938
  val argnames = Name.variant_list names
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   939
        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   940
  val args = map Free (argnames ~~ Ts)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   941
  val (inargs, outargs) = split_mode mode args
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   942
  val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   943
  val t = fold_rev lambda args r 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   944
in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   945
  (t, argnames @ names)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   946
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   947
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   948
fun create_intro_elim_rule nparams mode defthm mode_id funT pred thy =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   949
let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   950
  val Ts = binder_types (fastype_of pred)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   951
  val funtrm = Const (mode_id, funT)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   952
  val argnames = Name.variant_list []
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   953
        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   954
  val (Ts1, Ts2) = chop nparams Ts;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   955
  val Ts1' = map2 (param_funT_of (PredicateCompFuns.compfuns)) Ts1 (fst mode)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   956
  val args = map Free (argnames ~~ (Ts1' @ Ts2))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   957
  val (params, io_args) = chop nparams args
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   958
  val (inargs, outargs) = split_mode (snd mode) io_args
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   959
  val param_names = Name.variant_list argnames
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   960
    (map (fn i => "p" ^ string_of_int i) (1 upto nparams))
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   961
  val param_vs = map Free (param_names ~~ Ts1)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   962
  val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ (fst mode)) []
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   963
  val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ io_args))
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   964
  val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ io_args))
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   965
  val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   966
  val funargs = params @ inargs
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   967
  val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   968
                  if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   969
  val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   970
                   mk_tuple outargs))
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   971
  val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
   972
  val _ = tracing (Syntax.string_of_term_global thy introtrm) 
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   973
  val simprules = [defthm, @{thm eval_pred},
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   974
                   @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}]
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   975
  val unfolddef_tac = (Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1)
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   976
  val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   977
  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   978
  val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
   979
  val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   980
in 
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
   981
  (introthm, elimthm)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   982
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   983
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   984
fun create_constname_of_mode thy name mode = 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   985
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   986
    fun string_of_mode mode = if null mode then "0"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   987
      else space_implode "_" (map string_of_int mode)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   988
    fun string_of_HOmode m s = case m of NONE => s | SOME mode => s ^ "__" ^ (string_of_mode mode)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   989
    val HOmode = fold string_of_HOmode (fst mode) ""
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   990
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   991
    (Sign.full_bname thy (Long_Name.base_name name)) ^
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   992
      (if HOmode = "" then "_" else HOmode ^ "___") ^ (string_of_mode (snd mode))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   993
  end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   994
  
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   995
fun create_definitions preds nparams (name, modes) thy =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   996
  let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   997
    val _ = tracing "create definitions"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
   998
    val compfuns = PredicateCompFuns.compfuns
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   999
    val T = AList.lookup (op =) preds name |> the
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1000
    fun create_definition mode thy = let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1001
      val mode_cname = create_constname_of_mode thy name mode
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1002
      val mode_cbasename = Long_Name.base_name mode_cname
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1003
      val Ts = binder_types T;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1004
      val (Ts1, Ts2) = chop nparams Ts;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1005
      val Ts1' = map2 (param_funT_of compfuns) Ts1 (fst mode)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1006
      val (Us1, Us2) = split_mode (snd mode) Ts2;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1007
      val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1008
      val names = Name.variant_list []
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1009
        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1010
      val xs = map Free (names ~~ (Ts1' @ Ts2));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1011
      val (xparams, xargs) = chop nparams xs;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1012
      val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ (fst mode)) names
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1013
      val (xins, xouts) = split_mode (snd mode) xargs;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1014
      fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1015
       | mk_split_lambda [x] t = lambda x t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1016
       | mk_split_lambda xs t = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1017
         fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1018
           | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1019
         in mk_split_lambda' xs t end;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1020
      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts (list_comb (Const (name, T), xparams' @ xargs)))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1021
      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1022
      val def = Logic.mk_equals (lhs, predterm)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1023
      val ([definition], thy') = thy |>
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1024
        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1025
        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1026
      val (intro, elim) = create_intro_elim_rule nparams mode definition mode_cname funT (Const (name, T)) thy'
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1027
      in thy' |> add_predfun name mode (mode_cname, definition, intro, elim)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1028
        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1029
        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
31554
a6c6bf2751a0 added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents: 31550
diff changeset
  1030
        |> Theory.checkpoint
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1031
      end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1032
  in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1033
    fold create_definition modes thy
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1034
  end;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1035
  
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1036
(* TODO: use own theory datastructure for rpred *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1037
fun rpred_create_definitions preds nparams (name, modes) thy =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1038
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1039
    val T = AList.lookup (op =) preds name |> the
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1040
    fun create_definition mode thy =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1041
      let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1042
        val mode_cname = create_constname_of_mode thy name mode
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1043
        val Ts = binder_types T;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1044
        val (Ts1, Ts2) = chop nparams Ts;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1045
        val Ts1' = map2 (param_funT_of RPredCompFuns.compfuns) Ts1 (fst mode)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1046
        val (Us1, Us2) = split_mode (snd mode) Ts2;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1047
        val funT = (Ts1' @ Us1) ---> (RPredCompFuns.mk_rpredT (mk_tupleT Us2)) 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1048
      in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1049
        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1050
        |> add_predfun name mode (mode_cname, @{thm refl}, @{thm refl}, @{thm refl}) 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1051
      end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1052
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1053
    fold create_definition modes thy
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1054
  end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1055
  
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1056
(**************************************************************************************)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1057
(* Proving equivalence of term *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1058
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1059
fun is_Type (Type _) = true
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1060
  | is_Type _ = false
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1061
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1062
(* returns true if t is an application of an datatype constructor *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1063
(* which then consequently would be splitted *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1064
(* else false *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1065
fun is_constructor thy t =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1066
  if (is_Type (fastype_of t)) then
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
  1067
    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1068
      NONE => false
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1069
    | SOME info => (let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1070
      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1071
      val (c, _) = strip_comb t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1072
      in (case c of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1073
        Const (name, _) => name mem_string constr_consts
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1074
        | _ => false) end))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1075
  else false
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1076
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1077
(* MAJOR FIXME:  prove_params should be simple
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1078
 - different form of introrule for parameters ? *)
31549
f7379ea2c949 tuned; corrected exception handling
bulwahn
parents: 31541
diff changeset
  1079
fun prove_param thy modes (NONE, t) =
f7379ea2c949 tuned; corrected exception handling
bulwahn
parents: 31541
diff changeset
  1080
  all_tac 
f7379ea2c949 tuned; corrected exception handling
bulwahn
parents: 31541
diff changeset
  1081
| prove_param thy modes (m as SOME (Mode (mode, is, ms)), t) =
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1082
  REPEAT_DETERM (etac @{thm thin_rl} 1)
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1083
  THEN REPEAT_DETERM (rtac @{thm ext} 1)
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1084
  THEN (rtac @{thm iffI} 1)
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1085
  THEN print_tac "prove_param"
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1086
  (* proof in one direction *)
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1087
  THEN (atac 1)
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1088
  (* proof in the other direction *)
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1089
  THEN (atac 1)
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1090
  THEN print_tac "after prove_param"
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1091
(*  let
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1092
    val  (f, args) = strip_comb t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1093
    val (params, _) = chop (length ms) args
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1094
    val f_tac = case f of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1095
        Const (name, T) => simp_tac (HOL_basic_ss addsimps 
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1096
           (@{thm eval_pred}::(predfun_definition_of thy name mode)::
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1097
           @{thm "Product_Type.split_conv"}::[])) 1
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1098
      | Free _ => all_tac
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1099
      | Abs _ => error "TODO: implement here"
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1100
  in  
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1101
    print_tac "before simplification in prove_args:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1102
    THEN f_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1103
    THEN print_tac "after simplification in prove_args"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1104
    THEN (EVERY (map (prove_param thy modes) (ms ~~ params)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1105
    THEN (REPEAT_DETERM (atac 1))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1106
  end
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1107
*)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1108
fun prove_expr thy modes (SOME (Mode (mode, is, ms)), t, us) (premposition : int) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1109
  (case strip_comb t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1110
    (Const (name, T), args) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1111
      if AList.defined op = modes name then (let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1112
          val introrule = predfun_intro_of thy name mode
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1113
          (*val (in_args, out_args) = split_mode is us
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1114
          val (pred, rargs) = strip_comb (HOLogic.dest_Trueprop
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1115
            (hd (Logic.strip_imp_prems (prop_of introrule))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1116
          val nparams = length ms (* get_nparams thy (fst (dest_Const pred)) *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1117
          val (_, args) = chop nparams rargs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1118
          val subst = map (pairself (cterm_of thy)) (args ~~ us)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1119
          val inst_introrule = Drule.cterm_instantiate subst introrule*)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1120
         (* the next line is old and probably wrong *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1121
          val (args1, args2) = chop (length ms) args
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1122
        in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1123
        rtac @{thm bindI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1124
        THEN print_tac "before intro rule:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1125
        (* for the right assumption in first position *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1126
        THEN rotate_tac premposition 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1127
        THEN rtac introrule 1
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1128
        THEN print_tac "after intro rule"
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1129
        (* work with parameter arguments *)
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1130
        THEN (atac 1)
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1131
        THEN (print_tac "parameter goal")
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1132
        THEN (EVERY (map (prove_param thy modes) (ms ~~ args1)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1133
        THEN (REPEAT_DETERM (atac 1)) end)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1134
      else error "Prove expr if case not implemented"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1135
    | _ => rtac @{thm bindI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1136
           THEN atac 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1137
  | prove_expr _ _ _ _ =  error "Prove expr not implemented"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1138
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1139
fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1140
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1141
fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1142
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1143
fun prove_match thy (out_ts : term list) = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1144
  fun get_case_rewrite t =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1145
    if (is_constructor thy t) then let
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
  1146
      val case_rewrites = (#case_rewrites (Datatype.the_info thy
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1147
        ((fst o dest_Type o fastype_of) t)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1148
      in case_rewrites @ (flat (map get_case_rewrite (snd (strip_comb t)))) end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1149
    else []
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1150
  val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: (flat (map get_case_rewrite out_ts))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1151
(* replace TRY by determining if it necessary - are there equations when calling compile match? *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1152
in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1153
   (* make this simpset better! *)
31515
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1154
  asm_simp_tac (HOL_basic_ss' addsimps simprules) 1
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1155
  THEN print_tac "after prove_match:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1156
  THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1157
         THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1158
         THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1159
  THEN print_tac "after if simplification"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1160
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1161
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1162
(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1163
31515
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1164
fun prove_sidecond thy modes t =
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1165
  let
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1166
    fun preds_of t nameTs = case strip_comb t of 
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1167
      (f as Const (name, T), args) =>
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1168
        if AList.defined (op =) modes name then (name, T) :: nameTs
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1169
          else fold preds_of args nameTs
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1170
      | _ => nameTs
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1171
    val preds = preds_of t []
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1172
    val defs = map
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1173
      (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1174
        preds
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1175
  in 
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1176
    (* remove not_False_eq_True when simpset in prove_match is better *)
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1177
    simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1 
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1178
    (* need better control here! *)
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1179
  end
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1180
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1181
fun prove_clause thy nargs all_vs param_vs modes (iss, is) (ts, ps) = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1182
  val modes' = modes @ List.mapPartial
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1183
   (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1184
     (param_vs ~~ iss);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1185
  fun check_constrt ((names, eqs), t) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1186
      if is_constrt thy t then ((names, eqs), t) else
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1187
        let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1188
          val s = Name.variant names "x";
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1189
          val v = Free (s, fastype_of t)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1190
        in ((s::names, HOLogic.mk_eq (v, t)::eqs), v) end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1191
  
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1192
  val (in_ts, clause_out_ts) = split_mode is ts;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1193
  val ((all_vs', eqs), in_ts') =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1194
      (*FIXME*) Library.foldl_map check_constrt ((all_vs, []), in_ts);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1195
  fun prove_prems out_ts vs [] =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1196
    (prove_match thy out_ts)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1197
    THEN asm_simp_tac HOL_basic_ss' 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1198
    THEN print_tac "before the last rule of singleI:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1199
    THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1200
  | prove_prems out_ts vs rps =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1201
    let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1202
      val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1203
      val SOME (p, mode as SOME (Mode ((iss, js), _, param_modes))) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1204
        select_mode_prem thy modes' vs' rps;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1205
      val premposition = (find_index (equal p) ps) + nargs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1206
      val rps' = filter_out (equal p) rps;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1207
      val rest_tac = (case p of Prem (us, t) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1208
          let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1209
            val (in_ts, out_ts''') = split_mode js us
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1210
            val rec_tac = prove_prems out_ts''' vs' rps'
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1211
          in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1212
            print_tac "before clause:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1213
            THEN asm_simp_tac HOL_basic_ss 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1214
            THEN print_tac "before prove_expr:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1215
            THEN prove_expr thy modes (mode, t, us) premposition
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1216
            THEN print_tac "after prove_expr:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1217
            THEN rec_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1218
          end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1219
        | Negprem (us, t) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1220
          let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1221
            val (in_ts, out_ts''') = split_mode js us
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1222
            val rec_tac = prove_prems out_ts''' vs' rps'
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1223
            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1224
            val (_, params) = strip_comb t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1225
          in
31515
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1226
            rtac @{thm bindI} 1
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1227
            THEN (if (is_some name) then
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1228
                simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, js)]) 1
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1229
                THEN rtac @{thm not_predI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1230
                (* FIXME: work with parameter arguments *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1231
                THEN (EVERY (map (prove_param thy modes) (param_modes ~~ params)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1232
              else
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1233
                rtac @{thm not_predI'} 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1234
            THEN (REPEAT_DETERM (atac 1))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1235
            THEN rec_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1236
          end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1237
        | Sidecond t =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1238
         rtac @{thm bindI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1239
         THEN rtac @{thm if_predI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1240
         THEN print_tac "before sidecond:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1241
         THEN prove_sidecond thy modes t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1242
         THEN print_tac "after sidecond:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1243
         THEN prove_prems [] vs' rps')
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1244
    in (prove_match thy out_ts)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1245
        THEN rest_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1246
    end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1247
  val prems_tac = prove_prems in_ts' param_vs ps
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1248
in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1249
  rtac @{thm bindI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1250
  THEN rtac @{thm singleI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1251
  THEN prems_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1252
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1253
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1254
fun select_sup 1 1 = []
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1255
  | select_sup _ 1 = [rtac @{thm supI1}]
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1256
  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1257
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1258
fun prove_one_direction thy all_vs param_vs modes clauses ((pred, T), mode) = let
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31580
diff changeset
  1259
(*  val ind_result = Inductive.the_inductive (ProofContext.init thy) pred
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1260
  val index = find_index (fn s => s = pred) (#names (fst ind_result))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1261
  val (_, T) = dest_Const (nth (#preds (snd ind_result)) index) *)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1262
  val nargs = length (binder_types T) - nparams_of thy pred
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1263
  val pred_case_rule = singleton (ind_set_codegen_preproc thy)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1264
    (preprocess_elim thy nargs (the_elim_of thy pred))
31554
a6c6bf2751a0 added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents: 31550
diff changeset
  1265
  (* FIXME preprocessor |> Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm Predicate.memb_code}])*)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1266
in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1267
  REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1268
  THEN etac (predfun_elim_of thy pred mode) 1
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1269
  THEN etac pred_case_rule 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1270
  THEN (EVERY (map
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1271
         (fn i => EVERY' (select_sup (length clauses) i) i) 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1272
           (1 upto (length clauses))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1273
  THEN (EVERY (map (prove_clause thy nargs all_vs param_vs modes mode) clauses))
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1274
  THEN print_tac "proved one direction"
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1275
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1276
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1277
(** Proof in the other direction **)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1278
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1279
fun prove_match2 thy out_ts = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1280
  fun split_term_tac (Free _) = all_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1281
    | split_term_tac t =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1282
      if (is_constructor thy t) then let
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31723
diff changeset
  1283
        val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1284
        val num_of_constrs = length (#case_rewrites info)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1285
        (* special treatment of pairs -- because of fishing *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1286
        val split_rules = case (fst o dest_Type o fastype_of) t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1287
          "*" => [@{thm prod.split_asm}] 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1288
          | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm")
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1289
        val (_, ts) = strip_comb t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1290
      in
31515
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1291
        (Splitter.split_asm_tac split_rules 1)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1292
(*        THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1293
          THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1294
        THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1295
        THEN (EVERY (map split_term_tac ts))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1296
      end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1297
    else all_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1298
  in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1299
    split_term_tac (mk_tuple out_ts)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1300
    THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1301
  end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1302
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1303
(* VERY LARGE SIMILIRATIY to function prove_param 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1304
-- join both functions
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1305
*)
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
  1306
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1307
fun prove_param2 thy modes (NONE, t) = all_tac 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1308
  | prove_param2 thy modes (m as SOME (Mode (mode, is, ms)), t) = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1309
    val  (f, args) = strip_comb t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1310
    val (params, _) = chop (length ms) args
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1311
    val f_tac = case f of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1312
        Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1313
           (@{thm eval_pred}::(predfun_definition_of thy name mode)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1314
           :: @{thm "Product_Type.split_conv"}::[])) 1
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1315
      | Free _ => all_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1316
  in  
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1317
    print_tac "before simplification in prove_args:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1318
    THEN f_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1319
    THEN print_tac "after simplification in prove_args"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1320
    THEN (EVERY (map (prove_param2 thy modes) (ms ~~ params)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1321
  end
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
  1322
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1323
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1324
fun prove_expr2 thy modes (SOME (Mode (mode, is, ms)), t) = 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1325
  (case strip_comb t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1326
    (Const (name, T), args) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1327
      if AList.defined op = modes name then
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1328
        etac @{thm bindE} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1329
        THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1330
        THEN print_tac "prove_expr2-before"
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1331
        THEN (debug_tac (Syntax.string_of_term_global thy
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1332
          (prop_of (predfun_elim_of thy name mode))))
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1333
        THEN (etac (predfun_elim_of thy name mode) 1)
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1334
        THEN print_tac "prove_expr2"
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1335
        (* TODO -- FIXME: replace remove_last_goal*)
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
  1336
        (* THEN (EVERY (replicate (length args) (remove_last_goal thy))) *)
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
  1337
        THEN (EVERY (map (prove_param thy modes) (ms ~~ args)))
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1338
        THEN print_tac "finished prove_expr2"
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
  1339
      
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1340
      else error "Prove expr2 if case not implemented"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1341
    | _ => etac @{thm bindE} 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1342
  | prove_expr2 _ _ _ = error "Prove expr2 not implemented"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1343
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1344
fun prove_sidecond2 thy modes t = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1345
  fun preds_of t nameTs = case strip_comb t of 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1346
    (f as Const (name, T), args) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1347
      if AList.defined (op =) modes name then (name, T) :: nameTs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1348
        else fold preds_of args nameTs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1349
    | _ => nameTs
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1350
  val preds = preds_of t []
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1351
  val defs = map
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1352
    (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1353
      preds
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1354
  in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1355
   (* only simplify the one assumption *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1356
   full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1357
   (* need better control here! *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1358
   THEN print_tac "after sidecond2 simplification"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1359
   end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1360
  
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1361
fun prove_clause2 thy all_vs param_vs modes (iss, is) (ts, ps) pred i = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1362
  val modes' = modes @ List.mapPartial
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1363
   (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1364
     (param_vs ~~ iss);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1365
  fun check_constrt ((names, eqs), t) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1366
      if is_constrt thy t then ((names, eqs), t) else
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1367
        let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1368
          val s = Name.variant names "x";
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1369
          val v = Free (s, fastype_of t)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1370
        in ((s::names, HOLogic.mk_eq (v, t)::eqs), v) end;
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1371
  val pred_intro_rule = nth (intros_of thy pred) (i - 1)
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1372
    |> preprocess_intro thy
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1373
    |> (fn thm => hd (ind_set_codegen_preproc thy [thm]))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1374
    (* FIXME preprocess |> Simplifier.full_simplify (HOL_basic_ss addsimps [@ {thm Predicate.memb_code}]) *)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1375
  val (in_ts, clause_out_ts) = split_mode is ts;
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1376
  val ((all_vs', eqs), in_ts') =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1377
      (*FIXME*) Library.foldl_map check_constrt ((all_vs, []), in_ts);
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1378
  fun prove_prems2 out_ts vs [] =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1379
    print_tac "before prove_match2 - last call:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1380
    THEN prove_match2 thy out_ts
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1381
    THEN print_tac "after prove_match2 - last call:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1382
    THEN (etac @{thm singleE} 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1383
    THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1384
    THEN (asm_full_simp_tac HOL_basic_ss' 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1385
    THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1386
    THEN (asm_full_simp_tac HOL_basic_ss' 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1387
    THEN SOLVED (print_tac "state before applying intro rule:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1388
      THEN (rtac pred_intro_rule 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1389
      (* How to handle equality correctly? *)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1390
      THEN (print_tac "state before assumption matching")
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1391
      THEN (REPEAT (atac 1 ORELSE 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1392
         (CHANGED (asm_full_simp_tac HOL_basic_ss' 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1393
          THEN print_tac "state after simp_tac:"))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1394
  | prove_prems2 out_ts vs ps = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1395
      val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1396
      val SOME (p, mode as SOME (Mode ((iss, js), _, param_modes))) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1397
        select_mode_prem thy modes' vs' ps;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1398
      val ps' = filter_out (equal p) ps;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1399
      val rest_tac = (case p of Prem (us, t) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1400
          let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1401
            val (in_ts, out_ts''') = split_mode js us
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1402
            val rec_tac = prove_prems2 out_ts''' vs' ps'
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1403
          in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1404
            (prove_expr2 thy modes (mode, t)) THEN rec_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1405
          end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1406
        | Negprem (us, t) =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1407
          let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1408
            val (in_ts, out_ts''') = split_mode js us
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1409
            val rec_tac = prove_prems2 out_ts''' vs' ps'
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1410
            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1411
            val (_, params) = strip_comb t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1412
          in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1413
            print_tac "before neg prem 2"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1414
            THEN etac @{thm bindE} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1415
            THEN (if is_some name then
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1416
                full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, js)]) 1 
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1417
                THEN etac @{thm not_predE} 1
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1418
                THEN (EVERY (map (prove_param thy modes) (param_modes ~~ params)))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1419
              else
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1420
                etac @{thm not_predE'} 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1421
            THEN rec_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1422
          end 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1423
        | Sidecond t =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1424
            etac @{thm bindE} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1425
            THEN etac @{thm if_predE} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1426
            THEN prove_sidecond2 thy modes t 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1427
            THEN prove_prems2 [] vs' ps')
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1428
    in print_tac "before prove_match2:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1429
       THEN prove_match2 thy out_ts
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1430
       THEN print_tac "after prove_match2:"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1431
       THEN rest_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1432
    end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1433
  val prems_tac = prove_prems2 in_ts' param_vs ps 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1434
in
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1435
  print_tac "starting prove_clause2"
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1436
  THEN etac @{thm bindE} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1437
  THEN (etac @{thm singleE'} 1)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1438
  THEN (TRY (etac @{thm Pair_inject} 1))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1439
  THEN print_tac "after singleE':"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1440
  THEN prems_tac
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1441
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1442
 
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1443
fun prove_other_direction thy all_vs param_vs modes clauses (pred, mode) = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1444
  fun prove_clause (clause, i) =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1445
    (if i < length clauses then etac @{thm supE} 1 else all_tac)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1446
    THEN (prove_clause2 thy all_vs param_vs modes mode clause pred i)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1447
in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1448
  (DETERM (TRY (rtac @{thm unit.induct} 1)))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1449
   THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1450
   THEN (rtac (predfun_intro_of thy pred mode) 1)
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
  1451
   THEN (REPEAT_DETERM (rtac @{thm refl} 2))
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1452
   THEN (EVERY (map prove_clause (clauses ~~ (1 upto (length clauses)))))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1453
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1454
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1455
(** proof procedure **)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1456
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1457
fun prove_pred thy all_vs param_vs modes clauses (((pred, T), mode), t) = let
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1458
  val ctxt = ProofContext.init thy
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1459
  val clauses' = the (AList.lookup (op =) clauses pred)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1460
in
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1461
  Goal.prove ctxt (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) t []) [] t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1462
    (if !do_proofs then
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1463
      (fn _ =>
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1464
      rtac @{thm pred_iffI} 1
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1465
      THEN prove_one_direction thy all_vs param_vs modes clauses' ((pred, T), mode)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1466
      THEN print_tac "proved one direction"
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1467
      THEN prove_other_direction thy all_vs param_vs modes clauses' (pred, mode)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1468
      THEN print_tac "proved other direction")
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1469
     else (fn _ => mycheat_tac thy 1))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1470
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1471
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1472
fun prove_preds thy all_vs param_vs modes clauses pmts =
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1473
  map (prove_pred thy all_vs param_vs modes clauses) pmts
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1474
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1475
fun rpred_prove_preds thy pmts =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1476
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1477
    fun prove_pred (((pred, T), mode), t) =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1478
      let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1479
        val _ = Output.tracing ("prove_preds:" ^ Syntax.string_of_term_global thy t)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1480
      in SkipProof.make_thm thy t end
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1481
  in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1482
    map prove_pred pmts
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1483
  end
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1484
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1485
fun prepare_intrs thy prednames =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1486
  let
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1487
    val intrs = maps (intros_of thy) prednames
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1488
      |> map (Logic.unvarify o prop_of)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1489
    val nparams = nparams_of thy (hd prednames)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1490
    val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1491
    val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
32307
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
  1492
    val _ = Output.tracing ("extra_modes are: " ^
55166cd57a6d exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents: 32306
diff changeset
  1493
      cat_lines (map (fn (name, modes) => name ^ " has modes:" ^ (commas (map string_of_mode modes))) extra_modes))
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1494
    val _ $ u = Logic.strip_imp_concl (hd intrs);
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1495
    val params = List.take (snd (strip_comb u), nparams);
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1496
    val param_vs = maps term_vs params
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1497
    val all_vs = terms_vs intrs
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1498
    fun dest_prem t =
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1499
      (case strip_comb t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1500
        (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1501
      | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1502
          Prem (ts, t) => Negprem (ts, t)
31515
62fc203eed88 removed debug messages
bulwahn
parents: 31514
diff changeset
  1503
        | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) 
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1504
        | Sidecond t => Sidecond (c $ t))
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1505
      | (c as Const (s, _), ts) =>
31877
e3de75d3b898 exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents: 31876
diff changeset
  1506
        if is_registered thy s then
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1507
          let val (ts1, ts2) = chop (nparams_of thy s) ts
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1508
          in Prem (ts2, list_comb (c, ts1)) end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1509
        else Sidecond t
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1510
      | _ => Sidecond t)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1511
    fun add_clause intr (clauses, arities) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1512
    let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1513
      val _ $ t = Logic.strip_imp_concl intr;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1514
      val (Const (name, T), ts) = strip_comb t;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1515
      val (ts1, ts2) = chop nparams ts;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1516
      val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1517
      val (Ts, Us) = chop nparams (binder_types T)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1518
    in
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1519
      (AList.update op = (name, these (AList.lookup op = clauses name) @
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1520
        [(ts2, prems)]) clauses,
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1521
       AList.update op = (name, (map (fn U => (case strip_type U of
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1522
                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1523
               | _ => NONE)) Ts,
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1524
             length Us)) arities)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1525
    end;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1526
    val (clauses, arities) = fold add_clause intrs ([], []);
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1527
  in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end;
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1528
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1529
fun arrange kvs =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1530
  let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1531
    fun add (key, value) table =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1532
      AList.update op = (key, these (AList.lookup op = table key) @ [value]) table
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1533
  in fold add kvs [] end;
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1534
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1535
fun add_generators_to_clauses thy all_vs clauses =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1536
  let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1537
    val _ = Output.tracing ("all_vs:" ^ commas all_vs)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1538
  in  [clauses] end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1539
  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1540
(* main function *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1541
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1542
fun add_equations_of rpred prednames thy =
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1543
let
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1544
  val _ = tracing ("starting add_equations with " ^ commas prednames ^ "...")
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1545
  val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1546
    prepare_intrs thy prednames
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1547
  val _ = tracing "Infering modes..."
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1548
  (*
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1549
  val clauses_with_generators = add_generators_to_clauses thy all_vs clauses 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1550
  val modess = map (infer_modes thy extra_modes arities param_vs) clauses_with_generators
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1551
  fun print_modess (clauses, modes) =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1552
    let
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1553
      val _ = print_clausess thy clauses
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1554
      val _ = print_modes modes
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1555
    in
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1556
      ()
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1557
    end;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1558
  val _ = map print_modess (clauses_with_generators ~~ modess)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1559
  *)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1560
  val modes = infer_modes thy extra_modes arities param_vs clauses
31549
f7379ea2c949 tuned; corrected exception handling
bulwahn
parents: 31541
diff changeset
  1561
  val _ = print_modes modes
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1562
  val _ = tracing "Defining executable functions..."
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1563
  val thy' =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1564
    (if rpred then
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1565
      fold (rpred_create_definitions preds nparams) modes thy  
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1566
    else fold (create_definitions preds nparams) modes thy)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1567
    |> Theory.checkpoint
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1568
  val clauses' = map (fn (s, cls) => (s, (the (AList.lookup (op =) preds s), cls))) clauses
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1569
  val _ = tracing "Compiling equations..."
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1570
  val compfuns = if rpred then RPredCompFuns.compfuns else PredicateCompFuns.compfuns
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1571
  val ts = compile_preds thy' compfuns all_vs param_vs (extra_modes @ modes) clauses'
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1572
  val _ = map (Output.tracing o (Syntax.string_of_term_global thy')) (flat ts)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1573
  val pred_mode =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1574
    maps (fn (s, (T, _)) => map (pair (s, T)) ((the o AList.lookup (op =) modes) s)) clauses' 
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
  1575
  val _ = tracing "Proving equations..."
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1576
  val result_thms =
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1577
    if rpred then 
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1578
      rpred_prove_preds thy' (pred_mode ~~ (flat ts))
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1579
    else
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1580
      prove_preds thy' all_vs param_vs (extra_modes @ modes) clauses (pred_mode ~~ (flat ts))  
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1581
  val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1582
    [((Binding.qualify true (Long_Name.base_name name) (Binding.name "equation"), result_thms),
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1583
      [Attrib.attribute_i thy Code.add_default_eqn_attrib])] thy))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1584
    (arrange ((map (fn ((name, _), _) => name) pred_mode) ~~ result_thms)) thy'
31554
a6c6bf2751a0 added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents: 31550
diff changeset
  1585
    |> Theory.checkpoint
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1586
in
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1587
  thy''
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1588
end
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1589
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1590
(* generation of case rules from user-given introduction rules *)
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1591
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1592
fun mk_casesrule ctxt nparams introrules =
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1593
  let
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
  1594
    val intros = map (Logic.unvarify o prop_of) introrules
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1595
    val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1596
    val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1597
    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1598
    val (argnames, ctxt2) = Variable.variant_fixes
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1599
      (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1600
    val argvs = map2 (curry Free) argnames (map fastype_of args)
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1601
    fun mk_case intro = let
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1602
        val (_, (_, args)) = strip_intro_concl nparams intro
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1603
        val prems = Logic.strip_imp_prems intro
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1604
        val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args)
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1605
        val frees = (fold o fold_aterms)
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1606
          (fn t as Free _ =>
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1607
              if member (op aconv) params t then I else insert (op aconv) t
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1608
           | _ => I) (args @ prems) []
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1609
        in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1610
    val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1611
    val cases = map mk_case intros
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1612
  in Logic.list_implies (assm :: cases, prop) end;
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1613
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1614
fun add_equations name thy =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1615
  let
31554
a6c6bf2751a0 added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents: 31550
diff changeset
  1616
    val thy' = PredData.map (Graph.extend (dependencies_of thy) name) thy |> Theory.checkpoint;
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1617
    (*val preds = Graph.all_preds (PredData.get thy') [name] |> filter_out (has_elim thy') *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1618
    fun strong_conn_of gr keys =
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1619
      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1620
    val scc = strong_conn_of (PredData.get thy') [name]
31524
8abf99ab669c avoiding duplicate definitions of executable functions
bulwahn
parents: 31517
diff changeset
  1621
    val thy'' = fold_rev
8abf99ab669c avoiding duplicate definitions of executable functions
bulwahn
parents: 31517
diff changeset
  1622
      (fn preds => fn thy =>
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1623
        if forall (null o modes_of thy) preds then add_equations_of true preds thy else thy)
31554
a6c6bf2751a0 added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents: 31550
diff changeset
  1624
      scc thy' |> Theory.checkpoint
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1625
  in thy'' end
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1626
31876
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
  1627
  
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
  1628
fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
  1629
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
  1630
val code_pred_intros_attrib = attrib add_intro;
9ab571673059 added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents: 31784
diff changeset
  1631
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1632
(** user interface **)
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1633
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1634
local
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1635
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1636
(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1637
(* TODO: must create state to prove multiple cases *)
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1638
fun generic_code_pred prep_const raw_const lthy =
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1639
  let
31551
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
  1640
  
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1641
    val thy = ProofContext.theory_of lthy
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1642
    val const = prep_const thy raw_const
31551
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
  1643
    
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1644
    val lthy' = LocalTheory.theory (PredData.map (Graph.extend (dependencies_of thy) const)) lthy
31554
a6c6bf2751a0 added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents: 31550
diff changeset
  1645
      |> LocalTheory.checkpoint
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1646
    val thy' = ProofContext.theory_of lthy'
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1647
    val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
31551
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
  1648
    
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1649
    fun mk_cases const =
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1650
      let
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1651
        val nparams = nparams_of thy' const
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1652
        val intros = intros_of thy' const
31580
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1653
      in mk_casesrule lthy' nparams intros end  
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1654
    val cases_rules = map mk_cases preds
31580
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1655
    val cases =
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1656
      map (fn case_rule => RuleCases.Case {fixes = [],
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1657
        assumes = [("", Logic.strip_imp_prems case_rule)],
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1658
        binds = [], cases = []}) cases_rules
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1659
    val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1660
    val lthy'' = ProofContext.add_cases true case_env lthy'
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1661
    
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1662
    fun after_qed thms =
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31549
diff changeset
  1663
      LocalTheory.theory (fold set_elim (map the_single thms) #> add_equations const)
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1664
  in
31580
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31574
diff changeset
  1665
    Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1666
  end;
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1667
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1668
structure P = OuterParse
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1669
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1670
in
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1671
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1672
val code_pred = generic_code_pred (K I);
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31124
diff changeset
  1673
val code_pred_cmd = generic_code_pred Code.read_const
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1674
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1675
val setup = PredData.put (Graph.empty) #>
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1676
  Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro))
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1677
    "adding alternative introduction rules for code generation of inductive predicates"
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1678
(*  Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib)
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1679
    "adding alternative elimination rules for code generation of inductive predicates";
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1680
    *)
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1681
  (*FIXME name discrepancy in attribs and ML code*)
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1682
  (*FIXME intros should be better named intro*)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1683
  (*FIXME why distinguished attribute for cases?*)
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1684
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1685
val _ = OuterSyntax.local_theory_to_proof "code_pred"
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1686
  "prove equations for predicate specified by intro/elim rules"
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1687
  OuterKeyword.thy_goal (P.term_group >> code_pred_cmd)
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1688
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1689
end
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1690
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1691
(*FIXME
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1692
- Naming of auxiliary rules necessary?
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1693
- add default code equations P x y z = P_i_i_i x y z
31124
58bc773c60e2 marginally tuned
haftmann
parents: 31111
diff changeset
  1694
*)
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 31105
diff changeset
  1695
31169
f4c61cbea49d added predicate transformation function for code generation
bulwahn
parents: 31107
diff changeset
  1696
(* transformation for code generation *)
f4c61cbea49d added predicate transformation function for code generation
bulwahn
parents: 31107
diff changeset
  1697
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1698
val eval_ref = ref (NONE : (unit -> term Predicate.pred) option);
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1699
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1700
fun analyze_compr thy t_compr =
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1701
  let
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1702
    val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1703
      | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
32287
65d5c5b30747 cleaned up abstract tuple operations and named them consistently
haftmann
parents: 32091
diff changeset
  1704
    val (body, Ts, fp) = HOLogic.strip_splits split;
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1705
      (*FIXME former order of tuple positions must be restored*)
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1706
    val (pred as Const (name, T), all_args) = strip_comb body
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1707
    val (params, args) = chop (nparams_of thy name) all_args
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1708
    val user_mode = map_filter I (map_index
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1709
      (fn (i, t) => case t of Bound j => if j < length Ts then NONE
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1710
        else SOME (i+1) | _ => SOME (i+1)) args) (*FIXME dangling bounds should not occur*)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1711
    val (inargs, _) = split_mode user_mode args;
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1712
    val modes = filter (fn Mode (_, is, _) => is = user_mode)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31440
diff changeset
  1713
      (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1714
    val m = case modes
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1715
     of [] => error ("No mode possible for comprehension "
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1716
                ^ Syntax.string_of_term_global thy t_compr)
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1717
      | [m] => m
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1718
      | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1719
                ^ Syntax.string_of_term_global thy t_compr); m);
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents: 32307
diff changeset
  1720
    val t_eval = list_comb (compile_expr thy PredicateCompFuns.compfuns (all_modes_of thy) (SOME m, list_comb (pred, params)),
31217
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1721
      inargs)
c025f32afd4e experimental values command
haftmann
parents: 31177
diff changeset
  1722
  in t_eval end;
31169
f4c61cbea49d added predicate transformation function for code generation
bulwahn
parents: 31107
diff changeset
  1723
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1724
end;
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1725