src/HOL/Tools/Predicate_Compile/code_prolog.ML
author bulwahn
Thu, 26 Aug 2010 13:49:12 +0200
changeset 38789 d171840881fd
parent 38735 cb9031a9dccf
child 38790 499135eb21ec
permissions -rw-r--r--
added generation of predicates for size-limited enumeration of values
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Predicate_Compile/code_prolog.ML
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn, TU Muenchen
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     3
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     4
Prototype of an code generator for logic programming languages (a.k.a. Prolog)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     5
*)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     6
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     7
signature CODE_PROLOG =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
     8
sig
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
     9
  type code_options = {ensure_groundness : bool, limited_types : (typ * int) list}
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    10
  val options : code_options ref
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    11
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    12
  datatype arith_op = Plus | Minus
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
    13
  datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    14
    | Number of int | ArithOp of arith_op * prol_term list;
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    15
  datatype prem = Conj of prem list
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    16
    | Rel of string * prol_term list | NotRel of string * prol_term list
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    17
    | Eq of prol_term * prol_term | NotEq of prol_term * prol_term
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    18
    | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    19
    | Ground of string * typ;
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    20
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    21
  type clause = ((string * prol_term list) * prem);
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    22
  type logic_program = clause list;
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    23
  type constant_table = (string * string) list
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    24
    
38731
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
    25
  val generate : code_options -> Proof.context -> string -> (logic_program * constant_table)
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    26
  val write_program : logic_program -> string
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    27
  val run : logic_program -> string -> string list -> int option -> prol_term list list
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    28
38733
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
    29
  val quickcheck : Proof.context -> bool -> term -> int -> term list option * (bool list * bool)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
    30
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    31
  val trace : bool Unsynchronized.ref
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    32
end;
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    33
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    34
structure Code_Prolog : CODE_PROLOG =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    35
struct
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    36
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    37
(* diagnostic tracing *)
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    38
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    39
val trace = Unsynchronized.ref false
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    40
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
    41
fun tracing s = if !trace then Output.tracing s else () 
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    42
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    43
(* code generation options *)
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    44
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
    45
type code_options = {ensure_groundness : bool, limited_types : (typ * int) list}
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    46
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
    47
val options = Unsynchronized.ref {ensure_groundness = false, limited_types = []};
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    48
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    49
(* general string functions *)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    50
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    51
val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    52
val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    53
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    54
(* internal program representation *)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    55
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    56
datatype arith_op = Plus | Minus
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    57
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
    58
datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    59
  | Number of int | ArithOp of arith_op * prol_term list;
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    60
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    61
fun dest_Var (Var v) = v
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    62
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    63
fun add_vars (Var v) = insert (op =) v
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    64
  | add_vars (ArithOp (_, ts)) = fold add_vars ts
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    65
  | add_vars (AppF (_, ts)) = fold add_vars ts
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    66
  | add_vars _ = I
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    67
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    68
fun map_vars f (Var v) = Var (f v)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    69
  | map_vars f (ArithOp (opr, ts)) = ArithOp (opr, map (map_vars f) ts)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    70
  | map_vars f (AppF (fs, ts)) = AppF (fs, map (map_vars f) ts)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    71
  | map_vars f t = t
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    72
  
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
    73
fun maybe_AppF (c, []) = Cons c
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
    74
  | maybe_AppF (c, xs) = AppF (c, xs)
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
    75
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    76
fun is_Var (Var _) = true
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    77
  | is_Var _ = false
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    78
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    79
fun is_arith_term (Var _) = true
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    80
  | is_arith_term (Number _) = true
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    81
  | is_arith_term (ArithOp (_, operands)) = forall is_arith_term operands
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    82
  | is_arith_term _ = false
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    83
38081
8b02ce312893 removing pointless type information in internal prolog terms
bulwahn
parents: 38080
diff changeset
    84
fun string_of_prol_term (Var s) = "Var " ^ s
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
    85
  | string_of_prol_term (Cons s) = "Cons " ^ s
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
    86
  | string_of_prol_term (AppF (f, args)) = f ^ "(" ^ commas (map string_of_prol_term args) ^ ")" 
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
    87
  | string_of_prol_term (Number n) = "Number " ^ string_of_int n
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
    88
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    89
datatype prem = Conj of prem list
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    90
  | Rel of string * prol_term list | NotRel of string * prol_term list
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
    91
  | Eq of prol_term * prol_term | NotEq of prol_term * prol_term
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    92
  | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
    93
  | Ground of string * typ;
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    94
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
    95
fun dest_Rel (Rel (c, ts)) = (c, ts)
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    96
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    97
fun map_term_prem f (Conj prems) = Conj (map (map_term_prem f) prems)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    98
  | map_term_prem f (Rel (r, ts)) = Rel (r, map f ts)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
    99
  | map_term_prem f (NotRel (r, ts)) = NotRel (r, map f ts)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   100
  | map_term_prem f (Eq (l, r)) = Eq (f l, f r)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   101
  | map_term_prem f (NotEq (l, r)) = NotEq (f l, f r)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   102
  | map_term_prem f (ArithEq (l, r)) = ArithEq (f l, f r)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   103
  | map_term_prem f (NotArithEq (l, r)) = NotArithEq (f l, f r)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   104
  | map_term_prem f (Ground (v, T)) = Ground (dest_Var (f (Var v)), T)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   105
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   106
fun fold_prem_terms f (Conj prems) = fold (fold_prem_terms f) prems
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   107
  | fold_prem_terms f (Rel (_, ts)) = fold f ts
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   108
  | fold_prem_terms f (NotRel (_, ts)) = fold f ts
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   109
  | fold_prem_terms f (Eq (l, r)) = f l #> f r
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   110
  | fold_prem_terms f (NotEq (l, r)) = f l #> f r
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   111
  | fold_prem_terms f (ArithEq (l, r)) = f l #> f r
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   112
  | fold_prem_terms f (NotArithEq (l, r)) = f l #> f r
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   113
  | fold_prem_terms f (Ground (v, T)) = f (Var v)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   114
  
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   115
type clause = ((string * prol_term list) * prem);
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   116
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   117
type logic_program = clause list;
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   118
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   119
(* translation from introduction rules to internal representation *)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   120
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   121
(** constant table **)
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   122
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   123
type constant_table = (string * string) list
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   124
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   125
(* assuming no clashing *)
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   126
fun mk_constant_table consts =
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   127
  AList.make (first_lower o Long_Name.base_name) consts
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   128
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   129
fun declare_consts consts constant_table =
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   130
  fold (fn c => AList.update (op =) (c, first_lower (Long_Name.base_name c))) consts constant_table
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   131
  
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   132
fun translate_const constant_table c =
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   133
  case AList.lookup (op =) constant_table c of
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   134
    SOME c' => c'
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   135
  | NONE => error ("No such constant: " ^ c)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   136
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   137
fun inv_lookup _ [] _ = NONE
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   138
  | inv_lookup eq ((key, value)::xs) value' =
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   139
      if eq (value', value) then SOME key
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   140
      else inv_lookup eq xs value';
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   141
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   142
fun restore_const constant_table c =
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   143
  case inv_lookup (op =) constant_table c of
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   144
    SOME c' => c'
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   145
  | NONE => error ("No constant corresponding to "  ^ c)
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   146
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   147
(** translation of terms, literals, premises, and clauses **)
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   148
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   149
fun translate_arith_const @{const_name "Groups.plus_class.plus"} = SOME Plus
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   150
  | translate_arith_const @{const_name "Groups.minus_class.minus"} = SOME Minus
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   151
  | translate_arith_const _ = NONE
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   152
38734
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   153
fun mk_nat_term constant_table n =
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   154
  let
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   155
    val zero = translate_const constant_table @{const_name "Groups.zero_class.zero"}
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   156
    val Suc = translate_const constant_table @{const_name "Suc"}
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   157
  in funpow n (fn t => AppF (Suc, [t])) (Cons zero) end
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   158
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   159
fun translate_term ctxt constant_table t =
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   160
  case try HOLogic.dest_number t of
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   161
    SOME (@{typ "int"}, n) => Number n
38734
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   162
  | SOME (@{typ "nat"}, n) => mk_nat_term constant_table n
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   163
  | NONE =>
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   164
      (case strip_comb t of
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   165
        (Free (v, T), []) => Var v 
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   166
      | (Const (c, _), []) => Cons (translate_const constant_table c)
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   167
      | (Const (c, _), args) =>
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   168
        (case translate_arith_const c of
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   169
          SOME aop => ArithOp (aop, map (translate_term ctxt constant_table) args)
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   170
        | NONE =>                                                             
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   171
            AppF (translate_const constant_table c, map (translate_term ctxt constant_table) args))
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   172
      | _ => error ("illegal term for translation: " ^ Syntax.string_of_term ctxt t))
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   173
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   174
fun translate_literal ctxt constant_table t =
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   175
  case strip_comb t of
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   176
    (Const (@{const_name "op ="}, _), [l, r]) =>
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   177
      let
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   178
        val l' = translate_term ctxt constant_table l
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   179
        val r' = translate_term ctxt constant_table r
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   180
      in
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   181
        (if is_Var l' andalso is_arith_term r' andalso not (is_Var r') then ArithEq else Eq) (l', r')
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   182
      end
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   183
  | (Const (c, _), args) =>
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   184
      Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   185
  | _ => error ("illegal literal for translation: " ^ Syntax.string_of_term ctxt t)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   186
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   187
fun NegRel_of (Rel lit) = NotRel lit
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   188
  | NegRel_of (Eq eq) = NotEq eq
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   189
  | NegRel_of (ArithEq eq) = NotArithEq eq
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   190
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   191
fun mk_groundness_prems t = map Ground (Term.add_frees t [])
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   192
  
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   193
fun translate_prem options ctxt constant_table t =  
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   194
    case try HOLogic.dest_not t of
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   195
      SOME t =>
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   196
        if #ensure_groundness options then
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   197
          Conj (mk_groundness_prems t @ [NegRel_of (translate_literal ctxt constant_table t)])
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   198
        else
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   199
          NegRel_of (translate_literal ctxt constant_table t)
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   200
    | NONE => translate_literal ctxt constant_table t
38114
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   201
    
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   202
fun imp_prems_conv cv ct =
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   203
  case Thm.term_of ct of
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   204
    Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   205
  | _ => Conv.all_conv ct
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   206
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   207
fun Trueprop_conv cv ct =
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   208
  case Thm.term_of ct of
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   209
    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
38114
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   210
  | _ => raise Fail "Trueprop_conv"
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   211
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   212
fun preprocess_intro thy rule =
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   213
  Conv.fconv_rule
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   214
    (imp_prems_conv
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   215
      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq}))))
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   216
    (Thm.transfer thy rule)
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   217
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   218
fun translate_intros options ctxt gr const constant_table =
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   219
  let
38114
0f06e3cc04a6 adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
bulwahn
parents: 38113
diff changeset
   220
    val intros = map (preprocess_intro (ProofContext.theory_of ctxt)) (Graph.get_node gr const)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   221
    val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   222
    val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
38734
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
   223
      |> declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}]
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   224
    fun translate_intro intro =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   225
      let
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   226
        val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   227
        val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro)
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   228
        val prems' = Conj (map (translate_prem options ctxt' constant_table') prems)
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   229
        val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems')
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   230
      in clause end
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   231
  in (map translate_intro intros', constant_table') end
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   232
38731
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   233
val preprocess_options = Predicate_Compile_Aux.Options {
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   234
  expected_modes = NONE,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   235
  proposed_modes = NONE,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   236
  proposed_names = [],
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   237
  show_steps = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   238
  show_intermediate_results = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   239
  show_proof_trace = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   240
  show_modes = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   241
  show_mode_inference = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   242
  show_compilation = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   243
  show_caught_failures = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   244
  skip_proof = true,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   245
  no_topmost_reordering = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   246
  function_flattening = true,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   247
  specialise = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   248
  fail_safe_function_flattening = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   249
  no_higher_order_predicate = [],
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   250
  inductify = false,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   251
  detect_switches = true,
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   252
  compilation = Predicate_Compile_Aux.Pred
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   253
}
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   254
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   255
fun depending_preds_of (key, intros) =
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   256
  fold Term.add_const_names (map Thm.prop_of intros) []
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   257
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   258
fun add_edges edges_of key G =
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   259
  let
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   260
    fun extend' key (G, visited) = 
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   261
      case try (Graph.get_node G) key of
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   262
          SOME v =>
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   263
            let
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   264
              val new_edges = filter (fn k => is_some (try (Graph.get_node G) k)) (edges_of (key, v))
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   265
              val (G', visited') = fold extend'
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   266
                (subtract (op =) (key :: visited) new_edges) (G, key :: visited)
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   267
            in
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   268
              (fold (Graph.add_edge o (pair key)) new_edges G', visited')
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   269
            end
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   270
        | NONE => (G, visited)
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   271
  in
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   272
    fst (extend' key (G, []))
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   273
  end
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   274
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   275
fun generate options ctxt const =
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   276
  let 
38731
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   277
    fun strong_conn_of gr keys =
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   278
      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   279
    val gr = Predicate_Compile_Core.intros_graph_of ctxt
38731
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   280
    val gr' = add_edges depending_preds_of const gr
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38729
diff changeset
   281
    val scc = strong_conn_of gr' [const]
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   282
    val constant_table = mk_constant_table (flat scc)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   283
  in
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   284
    apfst flat (fold_map (translate_intros options ctxt gr) (flat scc) constant_table)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   285
  end
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   286
  
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   287
(* implementation for fully enumerating predicates and
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   288
  for size-limited predicates for enumerating the values of a datatype upto a specific size *)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   289
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   290
fun add_ground_typ (Conj prems) = fold add_ground_typ prems
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   291
  | add_ground_typ (Ground (_, T)) = insert (op =) T
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   292
  | add_ground_typ _ = I
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   293
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   294
fun mk_relname (Type (Tcon, Targs)) =
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   295
  first_lower (Long_Name.base_name Tcon) ^ space_implode "_" (map mk_relname Targs)
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   296
  | mk_relname _ = raise Fail "unexpected type"
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   297
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   298
fun mk_lim_relname T = "lim_" ^  mk_relname T
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   299
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   300
(* This is copied from "pat_completeness.ML" *)
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   301
fun inst_constrs_of thy (T as Type (name, _)) =
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   302
  map (fn (Cn,CT) =>
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   303
    Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   304
    (the (Datatype.get_constrs thy name))
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   305
  | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   306
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   307
fun is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   308
  
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   309
fun mk_ground_impl ctxt limited_types (T as Type (Tcon, Targs)) (seen, constant_table) =
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   310
  if member (op =) seen T then ([], (seen, constant_table))
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   311
  else
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   312
    let
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   313
      val (limited, size) = case AList.lookup (op =) limited_types T of
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   314
        SOME s => (true, s)
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   315
      | NONE => (false, 0)      
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   316
      val rel_name = (if limited then mk_lim_relname else mk_relname) T
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   317
      fun mk_impl (Const (constr_name, cT), recursive) (seen, constant_table) =
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   318
        let
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   319
          val constant_table' = declare_consts [constr_name] constant_table
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   320
          val Ts = binder_types cT
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   321
          val (rec_clauses, (seen', constant_table'')) =
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   322
            fold_map (mk_ground_impl ctxt limited_types) Ts (seen, constant_table')
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   323
          val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto (length Ts))
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   324
          val lim_var =
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   325
            if limited then
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   326
              if recursive then [AppF ("suc", [Var "Lim"])]              
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   327
              else [Var "Lim"]
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   328
            else [] 
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   329
          fun mk_prem v T' =
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   330
            if limited andalso T' = T then Rel (mk_lim_relname T', [Var "Lim", v])
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   331
            else Rel (mk_relname T', [v])
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   332
          val clause =
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   333
            ((rel_name, lim_var @ [maybe_AppF (translate_const constant_table'' constr_name, vars)]),
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   334
             Conj (map2 mk_prem vars Ts))
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   335
        in
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   336
          (clause :: flat rec_clauses, (seen', constant_table''))
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   337
        end
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   338
      val constrs = inst_constrs_of (ProofContext.theory_of ctxt) T
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   339
      val constrs' = (constrs ~~ map (is_recursive_constr T) constrs)
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   340
        |> (fn cs => filter_out snd cs @ filter snd cs)
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   341
      val (clauses, constant_table') =
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   342
        apfst flat (fold_map mk_impl constrs' (T :: seen, constant_table))
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   343
      val size_term = funpow size (fn t => AppF ("suc", [t])) (Cons "zero")
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   344
    in
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   345
      ((if limited then
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   346
        cons ((mk_relname T, [Var "x"]), Rel (mk_lim_relname T, [size_term, Var "x"]))
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   347
      else I) clauses, constant_table')
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   348
    end
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   349
 | mk_ground_impl ctxt _ T (seen, constant_table) =
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   350
   raise Fail ("unexpected type :" ^ Syntax.string_of_typ ctxt T)
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   351
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   352
fun replace_ground (Conj prems) = Conj (map replace_ground prems)
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   353
  | replace_ground (Ground (x, T)) =
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   354
    Rel (mk_relname T, [Var x])  
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   355
  | replace_ground p = p
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   356
  
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   357
fun add_ground_predicates ctxt limited_types (p, constant_table) =
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   358
  let
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   359
    val ground_typs = fold (add_ground_typ o snd) p []
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   360
    val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt limited_types) ground_typs ([], constant_table)
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   361
    val p' = map (apsnd replace_ground) p
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   362
  in
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   363
    ((flat grs) @ p', constant_table')
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   364
  end
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   365
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   366
  
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   367
(* rename variables to prolog-friendly names *)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   368
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   369
fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v))
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   370
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   371
fun rename_vars_prem renaming = map_term_prem (rename_vars_term renaming)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   372
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   373
fun dest_Char (Symbol.Char c) = c
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   374
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   375
fun is_prolog_conform v =
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   376
  forall (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s) (Symbol.explode v)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   377
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   378
fun mk_conform avoid v =
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   379
  let 
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   380
    val v' = space_implode "" (map (dest_Char o Symbol.decode)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   381
      (filter (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s)
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   382
        (Symbol.explode v)))
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   383
    val v' = if v' = "" then "var" else v'
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   384
  in Name.variant avoid (first_upper v') end
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   385
  
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   386
fun mk_renaming v renaming =
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   387
  (v, mk_conform (map snd renaming) v) :: renaming
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   388
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   389
fun rename_vars_clause ((rel, args), prem) =
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   390
  let
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   391
    val vars = fold_prem_terms add_vars prem (fold add_vars args [])
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   392
    val renaming = fold mk_renaming vars []
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   393
  in ((rel, map (rename_vars_term renaming) args), rename_vars_prem renaming prem) end
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   394
  
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   395
val rename_vars_program = map rename_vars_clause
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   396
  
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   397
(* code printer *)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   398
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   399
fun write_arith_op Plus = "+"
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   400
  | write_arith_op Minus = "-"
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   401
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   402
fun write_term (Var v) = v
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   403
  | write_term (Cons c) = c
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   404
  | write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   405
  | write_term (ArithOp (oper, [a1, a2])) = write_term a1 ^ " " ^ write_arith_op oper ^ " " ^ write_term a2
38112
cf08f4780938 adding numbers as basic term in prolog code generation
bulwahn
parents: 38082
diff changeset
   406
  | write_term (Number n) = string_of_int n
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   407
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   408
fun write_rel (pred, args) =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   409
  pred ^ "(" ^ space_implode ", " (map write_term args) ^ ")" 
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   410
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   411
fun write_prem (Conj prems) = space_implode ", " (map write_prem prems)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   412
  | write_prem (Rel p) = write_rel p  
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   413
  | write_prem (NotRel p) = "not(" ^ write_rel p ^ ")"
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   414
  | write_prem (Eq (l, r)) = write_term l ^ " = " ^ write_term r
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   415
  | write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r
38113
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   416
  | write_prem (ArithEq (l, r)) = write_term l ^ " is " ^ write_term r
81f08bbb3be7 adding basic arithmetic support for prolog code generation
bulwahn
parents: 38112
diff changeset
   417
  | write_prem (NotArithEq (l, r)) = write_term l ^ " =\\= " ^ write_term r
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   418
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   419
fun write_clause (head, prem) =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   420
  write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".")
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   421
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   422
fun write_program p =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   423
  cat_lines (map write_clause p) 
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   424
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   425
(** query templates **)
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   426
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   427
fun query_first rel vnames =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   428
  "eval :- once("  ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
38082
61280b97b761 adapting output for first solution
bulwahn
parents: 38081
diff changeset
   429
  "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
61280b97b761 adapting output for first solution
bulwahn
parents: 38081
diff changeset
   430
  "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   431
  
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   432
fun query_firstn n rel vnames =
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   433
  "eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   434
    rel ^ "(" ^ space_implode ", " vnames ^ "), Sols), writelist(Sols).\n" ^
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   435
    "writelist([]).\n" ^
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   436
    "writelist([(" ^ space_implode ", " vnames ^ ")|T]) :- " ^
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   437
    "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   438
    "\\n', [" ^ space_implode ", " vnames ^ "]), writelist(T).\n"
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   439
  
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   440
val prelude =
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   441
  "#!/usr/bin/swipl -q -t main -f\n\n" ^
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   442
  ":- use_module(library('dialect/ciao/aggregates')).\n" ^
38729
9c9d14827380 improving output of set comprehensions; adding style_check flags
bulwahn
parents: 38728
diff changeset
   443
  ":- style_check(-singleton).\n" ^
9c9d14827380 improving output of set comprehensions; adding style_check flags
bulwahn
parents: 38728
diff changeset
   444
  ":- style_check(-discontiguous).\n" ^ 	
9c9d14827380 improving output of set comprehensions; adding style_check flags
bulwahn
parents: 38728
diff changeset
   445
  ":- style_check(-atom).\n\n" ^
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   446
  "main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   447
  "main :- halt(1).\n"
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   448
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   449
(* parsing prolog solution *)
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   450
val scan_number =
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   451
  Scan.many1 Symbol.is_ascii_digit
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   452
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   453
val scan_atom =
38728
182b180e9804 improving ensure_groundness in prolog generation; added further example
bulwahn
parents: 38727
diff changeset
   454
  Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   455
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   456
val scan_var =
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   457
  Scan.many1
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   458
    (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   459
38076
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   460
val scan_ident =
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   461
  Scan.repeat (Scan.one
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   462
    (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s))
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   463
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   464
fun dest_Char (Symbol.Char s) = s
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   465
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   466
val string_of = concat o map (dest_Char o Symbol.decode)
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   467
38076
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   468
val is_atom_ident = forall Symbol.is_ascii_lower
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   469
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   470
val is_var_ident =
2a9c14d9d2ef correcting scanning
bulwahn
parents: 38075
diff changeset
   471
  forall (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   472
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   473
fun int_of_symbol_list xs = fold (fn x => fn s => s * 10 + (ord x - ord "0")) xs 0
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   474
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   475
fun scan_terms xs = (((scan_term --| $$ ",") ::: scan_terms)
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   476
  || (scan_term >> single)) xs
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   477
and scan_term xs =
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   478
  ((scan_number >> (Number o int_of_symbol_list))
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   479
  || (scan_var >> (Var o string_of))
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   480
  || ((scan_atom -- ($$ "(" |-- scan_terms --| $$ ")"))
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   481
    >> (fn (f, ts) => AppF (string_of f, ts)))
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   482
  || (scan_atom >> (Cons o string_of))) xs
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   483
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   484
val parse_term = fst o Scan.finite Symbol.stopper
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   485
    (Scan.error (!! (fn _ => raise Fail "parsing prolog output failed")) scan_term)
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   486
  o explode
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   487
  
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   488
fun parse_solutions sol =
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   489
  let
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   490
    fun dest_eq s = case space_explode "=" s of
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   491
        (l :: r :: []) => parse_term (unprefix " " r)
38078
2afb5f710b84 working on parser for prolog reponse
bulwahn
parents: 38077
diff changeset
   492
      | _ => raise Fail "unexpected equation in prolog output"
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   493
    fun parse_solution s = map dest_eq (space_explode ";" s)
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   494
  in
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   495
    map parse_solution (fst (split_last (space_explode "\n" sol)))
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   496
  end 
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   497
  
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   498
(* calling external interpreter and getting results *)
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   499
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   500
fun run p query_rel vnames nsols =
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   501
  let
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   502
    val cmd = Path.named_root
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   503
    val query = case nsols of NONE => query_first | SOME n => query_firstn n
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   504
    val p' = rename_vars_program p
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   505
    val _ = tracing "Renaming variable names..."
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   506
    val renaming = fold mk_renaming vnames [] 
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   507
    val vnames' = map (fn v => the (AList.lookup (op =) renaming v)) vnames
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   508
    val prog = prelude ^ query query_rel vnames' ^ write_program p'
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   509
    val _ = tracing ("Generated prolog program:\n" ^ prog)
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   510
    val prolog_file = File.tmp_path (Path.basic "prolog_file")
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   511
    val _ = File.write prolog_file prog
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   512
    val (solution, _) = bash_output ("/usr/local/bin/swipl -f " ^ File.shell_path prolog_file)
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   513
    val _ = tracing ("Prolog returned solution(s):\n" ^ solution)
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   514
    val tss = parse_solutions solution
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   515
  in
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   516
    tss
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   517
  end
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   518
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   519
(* values command *)
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   520
38081
8b02ce312893 removing pointless type information in internal prolog terms
bulwahn
parents: 38080
diff changeset
   521
fun restore_term ctxt constant_table (Var s, T) = Free (s, T)
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   522
  | restore_term ctxt constant_table (Number n, @{typ "int"}) = HOLogic.mk_number @{typ "int"} n
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   523
  | restore_term ctxt constant_table (Number n, _) = raise (Fail "unexpected type for number") 
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   524
  | restore_term ctxt constant_table (Cons s, T) = Const (restore_const constant_table s, T)
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   525
  | restore_term ctxt constant_table (AppF (f, args), T) =
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   526
    let
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   527
      val thy = ProofContext.theory_of ctxt
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   528
      val c = restore_const constant_table f
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   529
      val cT = Sign.the_const_type thy c
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   530
      val (argsT, resT) = strip_type cT
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   531
      val subst = Sign.typ_match thy (resT, T) Vartab.empty
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   532
      val argsT' = map (Envir.subst_type subst) argsT
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   533
    in
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   534
      list_comb (Const (c, Envir.subst_type subst cT),
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   535
        map (restore_term ctxt constant_table) (args ~~ argsT'))
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   536
    end
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   537
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   538
fun values ctxt soln t_compr =
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   539
  let
38727
c7f5f0b7dc7f adding very basic transformation to ensure groundness before negations
bulwahn
parents: 38558
diff changeset
   540
    val options = !options
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   541
    val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   542
      | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   543
    val (body, Ts, fp) = HOLogic.strip_psplits split;
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   544
    val output_names = Name.variant_list (Term.add_free_names body [])
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   545
      (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
38080
8c20eb9a388d cleaning example file; more natural ordering of variable names
bulwahn
parents: 38079
diff changeset
   546
    val output_frees = rev (map2 (curry Free) output_names Ts)
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   547
    val body = subst_bounds (output_frees, body)
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   548
    val (pred as Const (name, T), all_args) =
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   549
      case strip_comb body of
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   550
        (Const (name, T), all_args) => (Const (name, T), all_args)
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   551
      | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   552
    val vnames =
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   553
      case try (map (fst o dest_Free)) all_args of
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   554
        SOME vs => vs
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   555
      | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args))
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   556
    val _ = tracing "Preprocessing specification..."
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   557
    val T = Sign.the_const_type (ProofContext.theory_of ctxt) name
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   558
    val t = Const (name, T)
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   559
    val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   560
    val thy' = Predicate_Compile.preprocess preprocess_options t (ProofContext.theory_of ctxt')
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   561
    val ctxt'' = ProofContext.init_global thy'
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   562
    val _ = tracing "Generating prolog program..."
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   563
    val (p, constant_table) = generate options ctxt'' name
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   564
      |> (if #ensure_groundness options then
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   565
          add_ground_predicates ctxt'' (#limited_types options)
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   566
        else I)
38079
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   567
    val _ = tracing "Running prolog program..."
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   568
    val tss = run p (translate_const constant_table name) (map first_upper vnames) soln
7fb011dd51de improving translation to prolog; restoring terms from prolog output; adding tracing support
bulwahn
parents: 38078
diff changeset
   569
    val _ = tracing "Restoring terms..."
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   570
    val empty = Const("Orderings.bot_class.bot", fastype_of t_compr)
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   571
    fun mk_insert x S =
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   572
      Const (@{const_name "Set.insert"}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S 
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   573
    fun mk_set_compr in_insert [] xs =
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   574
       rev ((Free ("...", fastype_of t_compr)) ::
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   575
        (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs))
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   576
      | mk_set_compr in_insert (t :: ts) xs =
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   577
        let
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   578
          val frees = Term.add_frees t []
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   579
        in
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   580
          if null frees then
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   581
            mk_set_compr (t :: in_insert) ts xs
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   582
          else
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   583
            let
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   584
              val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt'' [t]) ("uu", fastype_of t)
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   585
              val set_compr =
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   586
                HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t))
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   587
                  frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"})))
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   588
            in
38729
9c9d14827380 improving output of set comprehensions; adding style_check flags
bulwahn
parents: 38728
diff changeset
   589
              mk_set_compr [] ts
9c9d14827380 improving output of set comprehensions; adding style_check flags
bulwahn
parents: 38728
diff changeset
   590
                (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs))  
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   591
            end
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   592
        end
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   593
  in
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   594
      foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr []
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   595
        (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt'' constant_table) (ts ~~ Ts))) tss) [])
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   596
  end
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   597
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   598
fun values_cmd print_modes soln raw_t state =
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   599
  let
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   600
    val ctxt = Toplevel.context_of state
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   601
    val t = Syntax.read_term ctxt raw_t
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   602
    val t' = values ctxt soln t
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   603
    val ty' = Term.type_of t'
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   604
    val ctxt' = Variable.auto_fixes t' ctxt
38115
987edb27f449 adding parsing of numbers; improving output of solution without free variables in prolog code generation
bulwahn
parents: 38114
diff changeset
   605
    val _ = tracing "Printing terms..."
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   606
    val p = Print_Mode.with_modes print_modes (fn () =>
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   607
      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   608
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   609
  in Pretty.writeln p end;
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   610
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   611
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   612
(* renewing the values command for Prolog queries *)
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   613
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   614
val opt_print_modes =
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   615
  Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   616
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   617
val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag
38077
46ff55ace18c querying for multiple solutions in values command for prolog execution
bulwahn
parents: 38076
diff changeset
   618
  (opt_print_modes -- Scan.optional (Parse.nat >> SOME) NONE -- Parse.term
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   619
   >> (fn ((print_modes, soln), t) => Toplevel.keep
38504
76965c356d2a removed Code_Prolog: modifies global environment setup non-conservatively
haftmann
parents: 38115
diff changeset
   620
        (values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*)
38075
3d5e2b7d1374 adding values command and parsing prolog output
bulwahn
parents: 38073
diff changeset
   621
38733
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   622
(* quickcheck generator *)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   623
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   624
(* FIXME: large copy of Predicate_Compile_Quickcheck - refactor out commons *)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   625
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   626
fun strip_imp_prems (Const(@{const_name "op -->"}, _) $ A $ B) = A :: strip_imp_prems B
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   627
  | strip_imp_prems _ = [];
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   628
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   629
fun strip_imp_concl (Const(@{const_name "op -->"}, _) $ A $ B) = strip_imp_concl B
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   630
  | strip_imp_concl A = A : term;
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   631
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   632
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   633
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   634
fun quickcheck ctxt report t size =
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   635
  let
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   636
    val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   637
    val thy = (ProofContext.theory_of ctxt')
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   638
    val (vs, t') = strip_abs t
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   639
    val vs' = Variable.variant_frees ctxt' [] vs
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   640
    val Ts = map snd vs'
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   641
    val t'' = subst_bounds (map Free (rev vs'), t')
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   642
    val (prems, concl) = strip_horn t''
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   643
    val constname = "quickcheck"
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   644
    val full_constname = Sign.full_bname thy constname
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   645
    val constT = Ts ---> @{typ bool}
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   646
    val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   647
    val const = Const (full_constname, constT)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   648
    val t = Logic.list_implies
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   649
      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   650
       HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   651
    val tac = fn _ => Skip_Proof.cheat_tac thy1
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   652
    val intro = Goal.prove (ProofContext.init_global thy1) (map fst vs') [] t tac
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   653
    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   654
    val thy3 = Predicate_Compile.preprocess preprocess_options const thy2
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   655
    val ctxt'' = ProofContext.init_global thy3
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   656
    val _ = tracing "Generating prolog program..."
38789
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   657
    val (p, constant_table) = generate {ensure_groundness = true, limited_types = []} ctxt'' full_constname
d171840881fd added generation of predicates for size-limited enumeration of values
bulwahn
parents: 38735
diff changeset
   658
      |> add_ground_predicates ctxt'' (#limited_types (!options))
38733
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   659
    val _ = tracing "Running prolog program..."
38735
cb9031a9dccf renaming variables to conform to prolog names
bulwahn
parents: 38734
diff changeset
   660
    val [ts] = run p (translate_const constant_table full_constname) (map fst vs')
38733
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   661
      (SOME 1)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   662
    val _ = tracing "Restoring terms..."
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   663
    val res = SOME (map (restore_term ctxt'' constant_table) (ts ~~ Ts))
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   664
    val empty_report = ([], false)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   665
  in
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   666
    (res, empty_report)
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38732
diff changeset
   667
  end; 
38732
3371dbc806ae moving preprocessing to values in prolog generation
bulwahn
parents: 38731
diff changeset
   668
38073
64062d56ad3c adding a mockup version for prolog code generation
bulwahn
parents:
diff changeset
   669
end;