src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Thu, 28 Jul 2011 11:43:45 +0200
changeset 44869 a2aa341bc658
parent 44853 05ff40b255eb
child 44959 3693baa6befb
permissions -rw-r--r--
tuning
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_provers.ML
     2     Author:     Fabian Immler, TU Muenchen
     3     Author:     Makarius
     4     Author:     Jasmin Blanchette, TU Muenchen
     5 
     6 Generic prover abstraction for Sledgehammer.
     7 *)
     8 
     9 signature SLEDGEHAMMER_PROVERS =
    10 sig
    11   type failure = ATP_Proof.failure
    12   type locality = ATP_Translate.locality
    13   type type_enc = ATP_Translate.type_enc
    14   type reconstructor = ATP_Reconstruct.reconstructor
    15   type play = ATP_Reconstruct.play
    16   type minimize_command = ATP_Reconstruct.minimize_command
    17   type relevance_fudge = Sledgehammer_Filter.relevance_fudge
    18 
    19   datatype mode = Auto_Try | Try | Normal | Minimize
    20 
    21   type params =
    22     {debug: bool,
    23      verbose: bool,
    24      overlord: bool,
    25      blocking: bool,
    26      provers: string list,
    27      type_enc: type_enc option,
    28      sound: bool,
    29      relevance_thresholds: real * real,
    30      max_relevant: int option,
    31      max_mono_iters: int,
    32      max_new_mono_instances: int,
    33      isar_proof: bool,
    34      isar_shrink_factor: int,
    35      slicing: bool,
    36      timeout: Time.time,
    37      preplay_timeout: Time.time,
    38      expect: string}
    39 
    40   datatype prover_fact =
    41     Untranslated_Fact of (string * locality) * thm |
    42     SMT_Weighted_Fact of (string * locality) * (int option * thm)
    43 
    44   type prover_problem =
    45     {state: Proof.state,
    46      goal: thm,
    47      subgoal: int,
    48      subgoal_count: int,
    49      facts: prover_fact list,
    50      smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
    51 
    52   type prover_result =
    53     {outcome: failure option,
    54      used_facts: (string * locality) list,
    55      run_time_in_msecs: int option,
    56      preplay: unit -> play,
    57      message: play -> string,
    58      message_tail: string}
    59 
    60   type prover =
    61     params -> (string -> minimize_command) -> prover_problem -> prover_result
    62 
    63   val concealedN : string
    64   val liftingN : string
    65   val combinatorsN : string
    66   val hybridN : string
    67   val lambdasN : string
    68   val smartN : string
    69   val dest_dir : string Config.T
    70   val problem_prefix : string Config.T
    71   val measure_run_time : bool Config.T
    72   val atp_lambda_translation : string Config.T
    73   val atp_readable_names : bool Config.T
    74   val smt_triggers : bool Config.T
    75   val smt_weights : bool Config.T
    76   val smt_weight_min_facts : int Config.T
    77   val smt_min_weight : int Config.T
    78   val smt_max_weight : int Config.T
    79   val smt_max_weight_index : int Config.T
    80   val smt_weight_curve : (int -> int) Unsynchronized.ref
    81   val smt_max_slices : int Config.T
    82   val smt_slice_fact_frac : real Config.T
    83   val smt_slice_time_frac : real Config.T
    84   val smt_slice_min_secs : int Config.T
    85   val das_tool : string
    86   val select_smt_solver : string -> Proof.context -> Proof.context
    87   val prover_name_for_reconstructor : reconstructor -> string option
    88   val is_metis_prover : string -> bool
    89   val is_atp : theory -> string -> bool
    90   val is_smt_prover : Proof.context -> string -> bool
    91   val is_unit_equational_atp : Proof.context -> string -> bool
    92   val is_prover_supported : Proof.context -> string -> bool
    93   val is_prover_installed : Proof.context -> string -> bool
    94   val default_max_relevant_for_prover : Proof.context -> bool -> string -> int
    95   val is_unit_equality : term -> bool
    96   val is_appropriate_prop_for_prover : Proof.context -> string -> term -> bool
    97   val is_built_in_const_for_prover :
    98     Proof.context -> string -> string * typ -> term list -> bool * term list
    99   val atp_relevance_fudge : relevance_fudge
   100   val smt_relevance_fudge : relevance_fudge
   101   val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
   102   val weight_smt_fact :
   103     Proof.context -> int -> ((string * locality) * thm) * int
   104     -> (string * locality) * (int option * thm)
   105   val untranslated_fact : prover_fact -> (string * locality) * thm
   106   val smt_weighted_fact :
   107     Proof.context -> int -> prover_fact * int
   108     -> (string * locality) * (int option * thm)
   109   val supported_provers : Proof.context -> unit
   110   val kill_provers : unit -> unit
   111   val running_provers : unit -> unit
   112   val messages : int option -> unit
   113   val filter_used_facts : ''a list -> (''a * 'b) list -> (''a * 'b) list
   114   val get_prover : Proof.context -> mode -> string -> prover
   115 end;
   116 
   117 structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
   118 struct
   119 
   120 open ATP_Util
   121 open ATP_Problem
   122 open ATP_Proof
   123 open ATP_Systems
   124 open ATP_Translate
   125 open ATP_Reconstruct
   126 open Sledgehammer_Util
   127 open Sledgehammer_Filter
   128 
   129 (** The Sledgehammer **)
   130 
   131 datatype mode = Auto_Try | Try | Normal | Minimize
   132 
   133 (* Identifier to distinguish Sledgehammer from other tools using
   134    "Async_Manager". *)
   135 val das_tool = "Sledgehammer"
   136 
   137 val metisN = Metis_Tactics.metisN
   138 val metis_full_typesN = metisN ^ "_" ^ Metis_Tactics.full_typesN
   139 val metis_no_typesN = metisN ^ "_" ^ Metis_Tactics.no_typesN
   140 
   141 val metis_prover_infos =
   142   [(metisN, Metis),
   143    (metis_full_typesN, Metis_Full_Types),
   144    (metis_no_typesN, Metis_No_Types)]
   145 
   146 val prover_name_for_reconstructor =
   147   AList.find (op =) metis_prover_infos #> try hd
   148 
   149 val metis_reconstructors = map snd metis_prover_infos
   150 
   151 val is_metis_prover = AList.defined (op =) metis_prover_infos
   152 val is_atp = member (op =) o supported_atps
   153 
   154 val select_smt_solver = Context.proof_map o SMT_Config.select_solver
   155 
   156 fun is_smt_prover ctxt name =
   157   member (op =) (SMT_Solver.available_solvers_of ctxt) name
   158 
   159 fun is_unit_equational_atp ctxt name =
   160   let val thy = Proof_Context.theory_of ctxt in
   161     case try (get_atp thy) name of
   162       SOME {formats, ...} => member (op =) formats CNF_UEQ
   163     | NONE => false
   164   end
   165 
   166 fun is_prover_supported ctxt name =
   167   let val thy = Proof_Context.theory_of ctxt in
   168     is_metis_prover name orelse is_atp thy name orelse is_smt_prover ctxt name
   169   end
   170 
   171 fun is_prover_installed ctxt =
   172   is_metis_prover orf is_smt_prover ctxt orf
   173   is_atp_installed (Proof_Context.theory_of ctxt)
   174 
   175 fun get_slices slicing slices =
   176   (0 upto length slices - 1) ~~ slices |> not slicing ? (List.last #> single)
   177 
   178 val metis_default_max_relevant = 20
   179 
   180 fun default_max_relevant_for_prover ctxt slicing name =
   181   let val thy = Proof_Context.theory_of ctxt in
   182     if is_metis_prover name then
   183       metis_default_max_relevant
   184     else if is_atp thy name then
   185       fold (Integer.max o #1 o snd o snd o snd)
   186            (get_slices slicing (#best_slices (get_atp thy name) ctxt)) 0
   187     else (* is_smt_prover ctxt name *)
   188       SMT_Solver.default_max_relevant ctxt name
   189   end
   190 
   191 fun is_if (@{const_name If}, _) = true
   192   | is_if _ = false
   193 
   194 (* Beware of "if and only if" (which is translated as such) and "If" (which is
   195    translated to conditional equations). *)
   196 fun is_good_unit_equality T t u =
   197   T <> @{typ bool} andalso not (exists (exists_Const is_if) [t, u])
   198 
   199 fun is_unit_equality (@{const Trueprop} $ t) = is_unit_equality t
   200   | is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) =
   201     is_unit_equality t
   202   | is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) =
   203     is_unit_equality t
   204   | is_unit_equality (Const (@{const_name "=="}, Type (_, [T, _])) $ t $ u) =
   205     is_good_unit_equality T t u
   206   | is_unit_equality (Const (@{const_name HOL.eq}, Type (_ , [T, _])) $ t $ u) =
   207     is_good_unit_equality T t u
   208   | is_unit_equality _ = false
   209 
   210 fun is_appropriate_prop_for_prover ctxt name =
   211   if is_unit_equational_atp ctxt name then is_unit_equality else K true
   212 
   213 fun is_built_in_const_for_prover ctxt name =
   214   if is_smt_prover ctxt name then
   215     let val ctxt = ctxt |> select_smt_solver name in
   216       fn x => fn ts =>
   217          if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
   218            (true, [])
   219          else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
   220            (true, ts)
   221          else
   222            (false, ts)
   223     end
   224   else
   225     fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts)
   226 
   227 (* FUDGE *)
   228 val atp_relevance_fudge =
   229   {local_const_multiplier = 1.5,
   230    worse_irrel_freq = 100.0,
   231    higher_order_irrel_weight = 1.05,
   232    abs_rel_weight = 0.5,
   233    abs_irrel_weight = 2.0,
   234    skolem_irrel_weight = 0.25,
   235    theory_const_rel_weight = 0.5,
   236    theory_const_irrel_weight = 0.25,
   237    chained_const_irrel_weight = 0.25,
   238    intro_bonus = 0.15,
   239    elim_bonus = 0.15,
   240    simp_bonus = 0.15,
   241    local_bonus = 0.55,
   242    assum_bonus = 1.05,
   243    chained_bonus = 1.5,
   244    max_imperfect = 11.5,
   245    max_imperfect_exp = 1.0,
   246    threshold_divisor = 2.0,
   247    ridiculous_threshold = 0.01}
   248 
   249 (* FUDGE (FIXME) *)
   250 val smt_relevance_fudge =
   251   {local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
   252    worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
   253    higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
   254    abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
   255    abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
   256    skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
   257    theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
   258    theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
   259    chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge,
   260    intro_bonus = #intro_bonus atp_relevance_fudge,
   261    elim_bonus = #elim_bonus atp_relevance_fudge,
   262    simp_bonus = #simp_bonus atp_relevance_fudge,
   263    local_bonus = #local_bonus atp_relevance_fudge,
   264    assum_bonus = #assum_bonus atp_relevance_fudge,
   265    chained_bonus = #chained_bonus atp_relevance_fudge,
   266    max_imperfect = #max_imperfect atp_relevance_fudge,
   267    max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
   268    threshold_divisor = #threshold_divisor atp_relevance_fudge,
   269    ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
   270 
   271 fun relevance_fudge_for_prover ctxt name =
   272   if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
   273 
   274 fun supported_provers ctxt =
   275   let
   276     val thy = Proof_Context.theory_of ctxt
   277     val (remote_provers, local_provers) =
   278       map fst metis_prover_infos @
   279       sort_strings (supported_atps thy) @
   280       sort_strings (SMT_Solver.available_solvers_of ctxt)
   281       |> List.partition (String.isPrefix remote_prefix)
   282   in
   283     Output.urgent_message ("Supported provers: " ^
   284                            commas (local_provers @ remote_provers) ^ ".")
   285   end
   286 
   287 fun kill_provers () = Async_Manager.kill_threads das_tool "prover"
   288 fun running_provers () = Async_Manager.running_threads das_tool "prover"
   289 val messages = Async_Manager.thread_messages das_tool "prover"
   290 
   291 (** problems, results, ATPs, etc. **)
   292 
   293 type params =
   294   {debug: bool,
   295    verbose: bool,
   296    overlord: bool,
   297    blocking: bool,
   298    provers: string list,
   299    type_enc: type_enc option,
   300    sound: bool,
   301    relevance_thresholds: real * real,
   302    max_relevant: int option,
   303    max_mono_iters: int,
   304    max_new_mono_instances: int,
   305    isar_proof: bool,
   306    isar_shrink_factor: int,
   307    slicing: bool,
   308    timeout: Time.time,
   309    preplay_timeout: Time.time,
   310    expect: string}
   311 
   312 datatype prover_fact =
   313   Untranslated_Fact of (string * locality) * thm |
   314   SMT_Weighted_Fact of (string * locality) * (int option * thm)
   315 
   316 type prover_problem =
   317   {state: Proof.state,
   318    goal: thm,
   319    subgoal: int,
   320    subgoal_count: int,
   321    facts: prover_fact list,
   322    smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
   323 
   324 type prover_result =
   325   {outcome: failure option,
   326    used_facts: (string * locality) list,
   327    run_time_in_msecs: int option,
   328    preplay: unit -> play,
   329    message: play -> string,
   330    message_tail: string}
   331 
   332 type prover =
   333   params -> (string -> minimize_command) -> prover_problem -> prover_result
   334 
   335 
   336 (* configuration attributes *)
   337 
   338 val concealedN = "concealed"
   339 val liftingN = "lifting"
   340 val combinatorsN = "combinators"
   341 val hybridN = "hybrid"
   342 val lambdasN = "lambdas"
   343 val smartN = "smart"
   344 
   345 (* Empty string means create files in Isabelle's temporary files directory. *)
   346 val dest_dir =
   347   Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
   348 val problem_prefix =
   349   Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
   350 val measure_run_time =
   351   Attrib.setup_config_bool @{binding sledgehammer_measure_run_time} (K false)
   352 
   353 val atp_lambda_translation =
   354   Attrib.setup_config_string @{binding sledgehammer_atp_lambda_translation}
   355                              (K smartN)
   356 (* In addition to being easier to read, readable names are often much shorter,
   357    especially if types are mangled in names. For these reason, they are enabled
   358    by default. *)
   359 val atp_readable_names =
   360   Attrib.setup_config_bool @{binding sledgehammer_atp_readable_names} (K true)
   361 
   362 val smt_triggers =
   363   Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
   364 val smt_weights =
   365   Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
   366 val smt_weight_min_facts =
   367   Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
   368 
   369 (* FUDGE *)
   370 val smt_min_weight =
   371   Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
   372 val smt_max_weight =
   373   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
   374 val smt_max_weight_index =
   375   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
   376 val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
   377 
   378 fun smt_fact_weight ctxt j num_facts =
   379   if Config.get ctxt smt_weights andalso
   380      num_facts >= Config.get ctxt smt_weight_min_facts then
   381     let
   382       val min = Config.get ctxt smt_min_weight
   383       val max = Config.get ctxt smt_max_weight
   384       val max_index = Config.get ctxt smt_max_weight_index
   385       val curve = !smt_weight_curve
   386     in
   387       SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
   388             div curve max_index)
   389     end
   390   else
   391     NONE
   392 
   393 fun weight_smt_fact ctxt num_facts ((info, th), j) =
   394   let val thy = Proof_Context.theory_of ctxt in
   395     (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
   396   end
   397 
   398 fun untranslated_fact (Untranslated_Fact p) = p
   399   | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
   400 fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
   401   | smt_weighted_fact ctxt num_facts (fact, j) =
   402     (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts
   403 
   404 fun overlord_file_location_for_prover prover =
   405   (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
   406 
   407 fun with_path cleanup after f path =
   408   Exn.capture f path
   409   |> tap (fn _ => cleanup path)
   410   |> Exn.release
   411   |> tap (after path)
   412 
   413 fun proof_banner mode name =
   414   case mode of
   415     Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
   416   | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
   417   | _ => "Try this"
   418 
   419 (* based on "Mirabelle.can_apply" and generalized *)
   420 fun timed_apply timeout tac state i =
   421   let
   422     val {context = ctxt, facts, goal} = Proof.goal state
   423     val full_tac = Method.insert_tac facts i THEN tac ctxt i
   424   in TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal end
   425 
   426 fun tac_for_reconstructor Metis =
   427     Metis_Tactics.metis_tac [Metis_Tactics.partial_type_enc]
   428   | tac_for_reconstructor Metis_Full_Types =
   429     Metis_Tactics.metis_tac Metis_Tactics.full_type_syss
   430   | tac_for_reconstructor Metis_No_Types =
   431     Metis_Tactics.metis_tac [Metis_Tactics.no_type_enc]
   432   | tac_for_reconstructor _ = raise Fail "unexpected reconstructor"
   433 
   434 fun timed_reconstructor reconstructor debug timeout ths =
   435   (Config.put Metis_Tactics.verbose debug
   436    #> (fn ctxt => tac_for_reconstructor reconstructor ctxt ths))
   437   |> timed_apply timeout
   438 
   439 fun filter_used_facts used = filter (member (op =) used o fst)
   440 
   441 fun play_one_line_proof debug timeout ths state i reconstructors =
   442   let
   443     fun play [] [] = Failed_to_Play (hd reconstructors)
   444       | play timed_outs [] = Trust_Playable (List.last timed_outs, SOME timeout)
   445       | play timed_out (reconstructor :: reconstructors) =
   446         let val timer = Timer.startRealTimer () in
   447           case timed_reconstructor reconstructor debug timeout ths state i of
   448             SOME (SOME _) => Played (reconstructor, Timer.checkRealTimer timer)
   449           | _ => play timed_out reconstructors
   450         end
   451         handle TimeLimit.TimeOut =>
   452                play (reconstructor :: timed_out) reconstructors
   453   in
   454     if timeout = Time.zeroTime then Trust_Playable (hd reconstructors, NONE)
   455     else play [] reconstructors
   456   end
   457 
   458 
   459 (* generic TPTP-based ATPs *)
   460 
   461 (* Too general means, positive equality literal with a variable X as one
   462    operand, when X does not occur properly in the other operand. This rules out
   463    clearly inconsistent facts such as X = a | X = b, though it by no means
   464    guarantees soundness. *)
   465 
   466 (* Unwanted equalities are those between a (bound or schematic) variable that
   467    does not properly occur in the second operand. *)
   468 val is_exhaustive_finite =
   469   let
   470     fun is_bad_equal (Var z) t =
   471         not (exists_subterm (fn Var z' => z = z' | _ => false) t)
   472       | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
   473       | is_bad_equal _ _ = false
   474     fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
   475     fun do_formula pos t =
   476       case (pos, t) of
   477         (_, @{const Trueprop} $ t1) => do_formula pos t1
   478       | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
   479         do_formula pos t'
   480       | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
   481         do_formula pos t'
   482       | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
   483         do_formula pos t'
   484       | (_, @{const "==>"} $ t1 $ t2) =>
   485         do_formula (not pos) t1 andalso
   486         (t2 = @{prop False} orelse do_formula pos t2)
   487       | (_, @{const HOL.implies} $ t1 $ t2) =>
   488         do_formula (not pos) t1 andalso
   489         (t2 = @{const False} orelse do_formula pos t2)
   490       | (_, @{const Not} $ t1) => do_formula (not pos) t1
   491       | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
   492       | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
   493       | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
   494       | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
   495       | _ => false
   496   in do_formula true end
   497 
   498 fun has_bound_or_var_of_type pred =
   499   exists_subterm (fn Var (_, T as Type _) => pred T
   500                    | Abs (_, T as Type _, _) => pred T
   501                    | _ => false)
   502 
   503 (* Facts are forbidden to contain variables of these types. The typical reason
   504    is that they lead to unsoundness. Note that "unit" satisfies numerous
   505    equations like "?x = ()". The resulting clauses will have no type constraint,
   506    yielding false proofs. Even "bool" leads to many unsound proofs, though only
   507    for higher-order problems. *)
   508 
   509 (* Facts containing variables of type "unit" or "bool" or of the form
   510    "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
   511    are omitted. *)
   512 fun is_dangerous_prop ctxt =
   513   transform_elim_prop
   514   #> (has_bound_or_var_of_type (is_type_surely_finite ctxt false) orf
   515       is_exhaustive_finite)
   516 
   517 fun int_opt_add (SOME m) (SOME n) = SOME (m + n)
   518   | int_opt_add _ _ = NONE
   519 
   520 (* Important messages are important but not so important that users want to see
   521    them each time. *)
   522 val atp_important_message_keep_quotient = 10
   523 
   524 fun choose_format_and_type_enc best_type_enc formats type_enc_opt =
   525   (case type_enc_opt of
   526      SOME ts => ts
   527    | NONE => type_enc_from_string best_type_enc)
   528   |> choose_format formats
   529 
   530 fun lift_lambdas ctxt type_enc =
   531   map (close_form o Envir.eta_contract) #> rpair ctxt
   532   #-> Lambda_Lifting.lift_lambdas
   533           (if polymorphism_of_type_enc type_enc = Polymorphic then
   534              SOME polymorphic_free_prefix
   535            else
   536              NONE)
   537           Lambda_Lifting.is_quantifier
   538   #> fst
   539 
   540 fun intentionalize_def (Const (@{const_name All}, _) $ Abs (_, _, t)) =
   541     intentionalize_def t
   542   | intentionalize_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
   543     let
   544       fun lam T t = Abs (Name.uu, T, t)
   545       val (head, args) = strip_comb t ||> rev
   546       val head_T = fastype_of head
   547       val n = length args
   548       val arg_Ts = head_T |> binder_types |> take n |> rev
   549       val u = u |> subst_atomic (args ~~ map Bound (0 upto n - 1))
   550     in HOLogic.eq_const head_T $ head $ fold lam arg_Ts u end
   551   | intentionalize_def t = t
   552 
   553 fun translate_atp_lambdas ctxt type_enc =
   554   Config.get ctxt atp_lambda_translation
   555   |> (fn trans =>
   556          if trans = smartN then
   557            if is_type_enc_higher_order type_enc then lambdasN else combinatorsN
   558          else if trans = lambdasN andalso
   559                  not (is_type_enc_higher_order type_enc) then
   560            error ("Lambda translation method incompatible with first-order \
   561                   \encoding.")
   562          else
   563            trans)
   564   |> (fn trans =>
   565          if trans = concealedN then
   566            lift_lambdas ctxt type_enc ##> K []
   567          else if trans = liftingN then
   568            lift_lambdas ctxt type_enc
   569          else if trans = combinatorsN then
   570            map (introduce_combinators ctxt) #> rpair []
   571          else if trans = hybridN then
   572            lift_lambdas ctxt type_enc
   573            ##> maps (fn t => [t, introduce_combinators ctxt
   574                                      (intentionalize_def t)])
   575          else if trans = lambdasN then
   576            map (Envir.eta_contract) #> rpair []
   577          else
   578            error ("Unknown lambda translation method: " ^ quote trans ^ "."))
   579 
   580 val metis_minimize_max_time = seconds 2.0
   581 
   582 fun choose_minimize_command minimize_command name preplay =
   583   (case preplay of
   584      Played (reconstructor, time) =>
   585      if Time.<= (time, metis_minimize_max_time) then
   586        prover_name_for_reconstructor reconstructor |> the_default name
   587      else
   588        name
   589    | _ => name)
   590   |> minimize_command
   591 
   592 fun repair_monomorph_context max_iters max_new_instances =
   593   Config.put Monomorph.max_rounds max_iters
   594   #> Config.put Monomorph.max_new_instances max_new_instances
   595   #> Config.put Monomorph.keep_partial_instances false
   596 
   597 (* Give the ATPs some slack before interrupting them the hard way. "z3_atp" on
   598    Linux appears to be the only ATP that does not honor its time limit. *)
   599 val atp_timeout_slack = seconds 1.0
   600 
   601 fun run_atp mode name
   602         ({exec, required_execs, arguments, proof_delims, known_failures,
   603           conj_sym_kind, prem_kind, formats, best_slices, ...} : atp_config)
   604         ({debug, verbose, overlord, type_enc, sound, max_relevant,
   605           max_mono_iters, max_new_mono_instances, isar_proof,
   606           isar_shrink_factor, slicing, timeout, preplay_timeout, ...} : params)
   607         minimize_command
   608         ({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) =
   609   let
   610     val thy = Proof.theory_of state
   611     val ctxt = Proof.context_of state
   612     val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
   613     val (dest_dir, problem_prefix) =
   614       if overlord then overlord_file_location_for_prover name
   615       else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
   616     val problem_file_name =
   617       Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
   618                   "_" ^ string_of_int subgoal)
   619     val problem_path_name =
   620       if dest_dir = "" then
   621         File.tmp_path problem_file_name
   622       else if File.exists (Path.explode dest_dir) then
   623         Path.append (Path.explode dest_dir) problem_file_name
   624       else
   625         error ("No such directory: " ^ quote dest_dir ^ ".")
   626     val measure_run_time = verbose orelse Config.get ctxt measure_run_time
   627     val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
   628     fun split_time s =
   629       let
   630         val split = String.tokens (fn c => str c = "\n")
   631         val (output, t) = s |> split |> split_last |> apfst cat_lines
   632         fun as_num f = f >> (fst o read_int)
   633         val num = as_num (Scan.many1 Symbol.is_ascii_digit)
   634         val digit = Scan.one Symbol.is_ascii_digit
   635         val num3 = as_num (digit ::: digit ::: (digit >> single))
   636         val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
   637         val as_time = Scan.read Symbol.stopper time o raw_explode
   638       in (output, as_time t) end
   639     fun run_on prob_file =
   640       case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
   641         (home_var, _) :: _ =>
   642         error ("The environment variable " ^ quote home_var ^ " is not set.")
   643       | [] =>
   644         if File.exists command then
   645           let
   646             (* If slicing is disabled, we expand the last slice to fill the
   647                entire time available. *)
   648             val actual_slices = get_slices slicing (best_slices ctxt)
   649             val num_actual_slices = length actual_slices
   650             fun monomorphize_facts facts =
   651               let
   652                 val ctxt =
   653                   ctxt |> repair_monomorph_context max_mono_iters
   654                                                    max_new_mono_instances
   655                 (* pseudo-theorem involving the same constants as the subgoal *)
   656                 val subgoal_th =
   657                   Logic.list_implies (hyp_ts, concl_t)
   658                   |> Skip_Proof.make_thm thy
   659                 val rths =
   660                   facts |> chop (length facts div 4)
   661                         |>> map (pair 1 o snd)
   662                         ||> map (pair 2 o snd)
   663                         |> op @
   664                         |> cons (0, subgoal_th)
   665               in
   666                 Monomorph.monomorph atp_schematic_consts_of rths ctxt
   667                 |> fst |> tl
   668                 |> curry ListPair.zip (map fst facts)
   669                 |> maps (fn (name, rths) => map (pair name o snd) rths)
   670               end
   671             fun run_slice (slice, (time_frac, (complete,
   672                                     (best_max_relevant, best_type_enc, extra))))
   673                           time_left =
   674               let
   675                 val num_facts =
   676                   length facts |> is_none max_relevant
   677                                   ? Integer.min best_max_relevant
   678                 val (format, type_enc) =
   679                   choose_format_and_type_enc best_type_enc formats type_enc
   680                 val fairly_sound = is_type_enc_fairly_sound type_enc
   681                 val facts =
   682                   facts |> map untranslated_fact
   683                         |> not fairly_sound
   684                            ? filter_out (is_dangerous_prop ctxt o prop_of o snd)
   685                         |> take num_facts
   686                         |> polymorphism_of_type_enc type_enc <> Polymorphic
   687                            ? monomorphize_facts
   688                         |> map (apsnd prop_of)
   689                 val real_ms = Real.fromInt o Time.toMilliseconds
   690                 val slice_timeout =
   691                   ((real_ms time_left
   692                     |> (if slice < num_actual_slices - 1 then
   693                           curry Real.min (time_frac * real_ms timeout)
   694                         else
   695                           I))
   696                    * 0.001) |> seconds
   697                 val generous_slice_timeout =
   698                   Time.+ (slice_timeout, atp_timeout_slack)
   699                 val _ =
   700                   if debug then
   701                     quote name ^ " slice #" ^ string_of_int (slice + 1) ^
   702                     " with " ^ string_of_int num_facts ^ " fact" ^
   703                     plural_s num_facts ^ " for " ^
   704                     string_from_time slice_timeout ^ "..."
   705                     |> Output.urgent_message
   706                   else
   707                     ()
   708                 val (atp_problem, pool, conjecture_offset, facts_offset,
   709                      fact_names, typed_helpers, sym_tab) =
   710                   prepare_atp_problem ctxt format conj_sym_kind prem_kind
   711                       type_enc sound false (translate_atp_lambdas ctxt type_enc)
   712                       (Config.get ctxt atp_readable_names) true hyp_ts concl_t
   713                       facts
   714                 fun weights () = atp_problem_weights atp_problem
   715                 val full_proof = debug orelse isar_proof
   716                 val core =
   717                   File.shell_path command ^ " " ^
   718                   arguments ctxt full_proof extra slice_timeout weights ^ " " ^
   719                   File.shell_path prob_file
   720                 val command =
   721                   (if measure_run_time then
   722                      "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }"
   723                    else
   724                      "exec " ^ core) ^ " 2>&1"
   725                 val _ =
   726                   atp_problem
   727                   |> tptp_lines_for_atp_problem format
   728                   |> cons ("% " ^ command ^ "\n")
   729                   |> File.write_list prob_file
   730                 val conjecture_shape =
   731                   conjecture_offset + 1
   732                     upto conjecture_offset + length hyp_ts + 1
   733                   |> map single
   734                 val ((output, msecs), (atp_proof, outcome)) =
   735                   TimeLimit.timeLimit generous_slice_timeout Isabelle_System.bash_output command
   736                   |>> (if overlord then
   737                          prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
   738                        else
   739                          I)
   740                   |> fst
   741                   |> (if measure_run_time then split_time else rpair NONE)
   742                   |> (fn accum as (output, _) =>
   743                          (accum,
   744                           extract_tstplike_proof_and_outcome verbose complete
   745                               proof_delims known_failures output
   746                           |>> atp_proof_from_tstplike_proof atp_problem output
   747                           handle UNRECOGNIZED_ATP_PROOF () =>
   748                                  ([], SOME ProofIncomplete)))
   749                   handle TimeLimit.TimeOut =>
   750                          (("", SOME (Time.toMilliseconds slice_timeout)),
   751                           ([], SOME TimedOut))
   752                 val outcome =
   753                   case outcome of
   754                     NONE =>
   755                     used_facts_in_unsound_atp_proof ctxt
   756                         conjecture_shape facts_offset fact_names atp_proof
   757                     |> Option.map (fn facts =>
   758                            UnsoundProof
   759                                (if is_type_enc_virtually_sound type_enc then
   760                                   SOME sound
   761                                 else
   762                                   NONE, facts |> sort string_ord))
   763                   | _ => outcome
   764               in
   765                 ((pool, conjecture_shape, facts_offset, fact_names,
   766                   typed_helpers, sym_tab),
   767                  (output, msecs, atp_proof, outcome))
   768               end
   769             val timer = Timer.startRealTimer ()
   770             fun maybe_run_slice slice (result as (_, (_, msecs0, _, SOME _))) =
   771                 let
   772                   val time_left = Time.- (timeout, Timer.checkRealTimer timer)
   773                 in
   774                   if Time.<= (time_left, Time.zeroTime) then
   775                     result
   776                   else
   777                     (run_slice slice time_left
   778                      |> (fn (stuff, (output, msecs, atp_proof, outcome)) =>
   779                             (stuff, (output, int_opt_add msecs0 msecs,
   780                                      atp_proof, outcome))))
   781                 end
   782               | maybe_run_slice _ result = result
   783           in
   784             ((Symtab.empty, [], 0, Vector.fromList [], [], Symtab.empty),
   785              ("", SOME 0, [], SOME InternalError))
   786             |> fold maybe_run_slice actual_slices
   787           end
   788         else
   789           error ("Bad executable: " ^ Path.print command)
   790 
   791     (* If the problem file has not been exported, remove it; otherwise, export
   792        the proof file too. *)
   793     fun cleanup prob_file =
   794       if dest_dir = "" then try File.rm prob_file else NONE
   795     fun export prob_file (_, (output, _, _, _)) =
   796       if dest_dir = "" then
   797         ()
   798       else
   799         File.write (Path.explode (Path.implode prob_file ^ "_proof")) output
   800     val ((pool, conjecture_shape, facts_offset, fact_names, typed_helpers,
   801           sym_tab),
   802          (output, msecs, atp_proof, outcome)) =
   803       with_path cleanup export run_on problem_path_name
   804     val important_message =
   805       if mode = Normal andalso
   806          random_range 0 (atp_important_message_keep_quotient - 1) = 0 then
   807         extract_important_message output
   808       else
   809         ""
   810     val (used_facts, preplay, message, message_tail) =
   811       case outcome of
   812         NONE =>
   813         let
   814           val used_facts =
   815             used_facts_in_atp_proof ctxt facts_offset fact_names atp_proof
   816         in
   817           (used_facts,
   818            fn () =>
   819               let
   820                 val used_ths =
   821                   facts |> map untranslated_fact |> filter_used_facts used_facts
   822                         |> map snd
   823               in
   824                 (if mode = Minimize andalso
   825                     Time.> (preplay_timeout, Time.zeroTime) then
   826                    Output.urgent_message "Preplaying proof..."
   827                  else
   828                    ());
   829                 play_one_line_proof debug preplay_timeout used_ths state subgoal
   830                                     metis_reconstructors
   831               end,
   832            fn preplay =>
   833               let
   834                 val full_types = uses_typed_helpers typed_helpers atp_proof
   835                 val isar_params =
   836                   (debug, full_types, isar_shrink_factor, pool,
   837                    conjecture_shape, facts_offset, fact_names, sym_tab,
   838                    atp_proof, goal)
   839                 val one_line_params =
   840                   (preplay, proof_banner mode name, used_facts,
   841                    choose_minimize_command minimize_command name preplay,
   842                    subgoal, subgoal_count)
   843               in proof_text ctxt isar_proof isar_params one_line_params end,
   844            (if verbose then
   845               "\nATP real CPU time: " ^
   846               string_from_time (Time.fromMilliseconds (the msecs)) ^ "."
   847             else
   848               "") ^
   849            (if important_message <> "" then
   850               "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^
   851               important_message
   852             else
   853               ""))
   854         end
   855       | SOME failure =>
   856         ([], K (Failed_to_Play Metis), fn _ => string_for_failure failure, "")
   857   in
   858     {outcome = outcome, used_facts = used_facts, run_time_in_msecs = msecs,
   859      preplay = preplay, message = message, message_tail = message_tail}
   860   end
   861 
   862 (* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
   863    these are sorted out properly in the SMT module, we have to interpret these
   864    ourselves. *)
   865 val remote_smt_failures =
   866   [(2, NoLibwwwPerl),
   867    (22, CantConnect)]
   868 val z3_failures =
   869   [(101, OutOfResources),
   870    (103, MalformedInput),
   871    (110, MalformedInput)]
   872 val unix_failures =
   873   [(139, Crashed)]
   874 val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
   875 
   876 fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
   877     if is_real_cex then Unprovable else GaveUp
   878   | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
   879   | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
   880     (case AList.lookup (op =) smt_failures code of
   881        SOME failure => failure
   882      | NONE => UnknownError ("Abnormal termination with exit code " ^
   883                              string_of_int code ^ "."))
   884   | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
   885   | failure_from_smt_failure (SMT_Failure.Other_Failure msg) =
   886     UnknownError msg
   887 
   888 (* FUDGE *)
   889 val smt_max_slices =
   890   Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
   891 val smt_slice_fact_frac =
   892   Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
   893 val smt_slice_time_frac =
   894   Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
   895 val smt_slice_min_secs =
   896   Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
   897 
   898 fun smt_filter_loop ctxt name
   899                     ({debug, verbose, overlord, max_mono_iters,
   900                       max_new_mono_instances, timeout, slicing, ...} : params)
   901                     state i smt_filter =
   902   let
   903     val max_slices = if slicing then Config.get ctxt smt_max_slices else 1
   904     val repair_context =
   905       select_smt_solver name
   906       #> (if overlord then
   907             Config.put SMT_Config.debug_files
   908                        (overlord_file_location_for_prover name
   909                         |> (fn (path, name) => path ^ "/" ^ name))
   910           else
   911             I)
   912       #> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers)
   913     val state = state |> Proof.map_context repair_context
   914     fun do_slice timeout slice outcome0 time_so_far facts =
   915       let
   916         val timer = Timer.startRealTimer ()
   917         val state =
   918           state |> Proof.map_context
   919                        (repair_monomorph_context max_mono_iters
   920                                                  max_new_mono_instances)
   921         val ms = timeout |> Time.toMilliseconds
   922         val slice_timeout =
   923           if slice < max_slices then
   924             Int.min (ms,
   925                 Int.max (1000 * Config.get ctxt smt_slice_min_secs,
   926                     Real.ceil (Config.get ctxt smt_slice_time_frac
   927                                * Real.fromInt ms)))
   928             |> Time.fromMilliseconds
   929           else
   930             timeout
   931         val num_facts = length facts
   932         val _ =
   933           if debug then
   934             quote name ^ " slice " ^ string_of_int slice ^ " with " ^
   935             string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^
   936             string_from_time slice_timeout ^ "..."
   937             |> Output.urgent_message
   938           else
   939             ()
   940         val birth = Timer.checkRealTimer timer
   941         val _ =
   942           if debug then Output.urgent_message "Invoking SMT solver..." else ()
   943         val (outcome, used_facts) =
   944           (case (slice, smt_filter) of
   945              (1, SOME head) => head |> apsnd (apsnd repair_context)
   946            | _ => SMT_Solver.smt_filter_preprocess state facts i)
   947           |> SMT_Solver.smt_filter_apply slice_timeout
   948           |> (fn {outcome, used_facts} => (outcome, used_facts))
   949           handle exn => if Exn.is_interrupt exn then
   950                           reraise exn
   951                         else
   952                           (ML_Compiler.exn_message exn
   953                            |> SMT_Failure.Other_Failure |> SOME, [])
   954         val death = Timer.checkRealTimer timer
   955         val outcome0 = if is_none outcome0 then SOME outcome else outcome0
   956         val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
   957         val too_many_facts_perhaps =
   958           case outcome of
   959             NONE => false
   960           | SOME (SMT_Failure.Counterexample _) => false
   961           | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
   962           | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
   963           | SOME SMT_Failure.Out_Of_Memory => true
   964           | SOME (SMT_Failure.Other_Failure _) => true
   965         val timeout = Time.- (timeout, Timer.checkRealTimer timer)
   966       in
   967         if too_many_facts_perhaps andalso slice < max_slices andalso
   968            num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
   969           let
   970             val new_num_facts =
   971               Real.ceil (Config.get ctxt smt_slice_fact_frac
   972                          * Real.fromInt num_facts)
   973             val _ =
   974               if verbose andalso is_some outcome then
   975                 quote name ^ " invoked with " ^ string_of_int num_facts ^
   976                 " fact" ^ plural_s num_facts ^ ": " ^
   977                 string_for_failure (failure_from_smt_failure (the outcome)) ^
   978                 " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^
   979                 plural_s new_num_facts ^ "..."
   980                 |> Output.urgent_message
   981               else
   982                 ()
   983           in
   984             facts |> take new_num_facts
   985                   |> do_slice timeout (slice + 1) outcome0 time_so_far
   986           end
   987         else
   988           {outcome = if is_none outcome then NONE else the outcome0,
   989            used_facts = used_facts,
   990            run_time_in_msecs = SOME (Time.toMilliseconds time_so_far)}
   991       end
   992   in do_slice timeout 1 NONE Time.zeroTime end
   993 
   994 fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
   995         minimize_command
   996         ({state, subgoal, subgoal_count, facts, smt_filter, ...}
   997          : prover_problem) =
   998   let
   999     val ctxt = Proof.context_of state
  1000     val num_facts = length facts
  1001     val facts = facts ~~ (0 upto num_facts - 1)
  1002                 |> map (smt_weighted_fact ctxt num_facts)
  1003     val {outcome, used_facts, run_time_in_msecs} =
  1004       smt_filter_loop ctxt name params state subgoal smt_filter facts
  1005     val (used_facts, used_ths) = used_facts |> ListPair.unzip
  1006     val outcome = outcome |> Option.map failure_from_smt_failure
  1007     val (preplay, message, message_tail) =
  1008       case outcome of
  1009         NONE =>
  1010         (fn () =>
  1011             let
  1012               fun smt_settings () =
  1013                 if name = SMT_Solver.solver_name_of ctxt then ""
  1014                 else "smt_solver = " ^ maybe_quote name
  1015             in
  1016               case play_one_line_proof debug preplay_timeout used_ths state
  1017                                        subgoal metis_reconstructors of
  1018                 p as Played _ => p
  1019               | _ => Trust_Playable (SMT (smt_settings ()), NONE)
  1020             end,
  1021          fn preplay =>
  1022             let
  1023               val one_line_params =
  1024                 (preplay, proof_banner mode name, used_facts,
  1025                  choose_minimize_command minimize_command name preplay,
  1026                  subgoal, subgoal_count)
  1027             in one_line_proof_text one_line_params end,
  1028          if verbose then
  1029            "\nSMT solver real CPU time: " ^
  1030            string_from_time (Time.fromMilliseconds
  1031                                  (the run_time_in_msecs)) ^ "."
  1032          else
  1033            "")
  1034       | SOME failure =>
  1035         (K (Failed_to_Play Metis), fn _ => string_for_failure failure, "")
  1036   in
  1037     {outcome = outcome, used_facts = used_facts,
  1038      run_time_in_msecs = run_time_in_msecs, preplay = preplay,
  1039      message = message, message_tail = message_tail}
  1040   end
  1041 
  1042 fun run_metis mode name ({debug, timeout, ...} : params) minimize_command
  1043               ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
  1044   let
  1045     val reconstructor =
  1046       case AList.lookup (op =) metis_prover_infos name of
  1047         SOME r => r
  1048       | NONE => raise Fail ("unknown Metis version: " ^ quote name)
  1049     val (used_facts, used_ths) =
  1050       facts |> map untranslated_fact |> ListPair.unzip
  1051   in
  1052     case play_one_line_proof debug timeout used_ths state subgoal
  1053                              [reconstructor] of
  1054       play as Played (_, time) =>
  1055       {outcome = NONE, used_facts = used_facts,
  1056        run_time_in_msecs = SOME (Time.toMilliseconds time),
  1057        preplay = K play,
  1058        message = fn play =>
  1059                     let
  1060                       val one_line_params =
  1061                          (play, proof_banner mode name, used_facts,
  1062                           minimize_command name, subgoal, subgoal_count)
  1063                     in one_line_proof_text one_line_params end,
  1064        message_tail = ""}
  1065     | play =>
  1066       let
  1067         val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut
  1068       in
  1069         {outcome = SOME failure, used_facts = [], run_time_in_msecs = NONE,
  1070          preplay = K play, message = fn _ => string_for_failure failure,
  1071          message_tail = ""}
  1072       end
  1073   end
  1074 
  1075 fun get_prover ctxt mode name =
  1076   let val thy = Proof_Context.theory_of ctxt in
  1077     if is_metis_prover name then run_metis mode name
  1078     else if is_atp thy name then run_atp mode name (get_atp thy name)
  1079     else if is_smt_prover ctxt name then run_smt_solver mode name
  1080     else error ("No such prover: " ^ name ^ ".")
  1081   end
  1082 
  1083 end;