src/HOL/Tools/ATP/atp_proof.ML
changeset 67022 49309fe530fd
parent 67021 41f1f8c4259b
child 68250 c45067867860
equal deleted inserted replaced
67021:41f1f8c4259b 67022:49309fe530fd
    89   val clean_up_atp_proof_dependencies : string atp_proof -> string atp_proof
    89   val clean_up_atp_proof_dependencies : string atp_proof -> string atp_proof
    90   val map_term_names_in_atp_proof : (string -> string) -> string atp_proof -> string atp_proof
    90   val map_term_names_in_atp_proof : (string -> string) -> string atp_proof -> string atp_proof
    91   val nasty_atp_proof : string Symtab.table -> string atp_proof -> string atp_proof
    91   val nasty_atp_proof : string Symtab.table -> string atp_proof -> string atp_proof
    92 
    92 
    93   val skip_term: string list -> string * string list
    93   val skip_term: string list -> string * string list
    94   val parse_thf0_formula :string list ->
    94   val parse_thf_formula :string list ->
    95     ('a, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, 'c) ATP_Problem.atp_formula *
    95     ('a, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, 'c) ATP_Problem.atp_formula *
    96     string list
    96     string list
    97   val dummy_atype : string ATP_Problem.atp_type
    97   val dummy_atype : string ATP_Problem.atp_type
    98   val role_of_tptp_string: string -> ATP_Problem.atp_formula_role
    98   val role_of_tptp_string: string -> ATP_Problem.atp_formula_role
    99   val parse_line: string -> ('a * string ATP_Problem.atp_problem_line list) list ->
    99   val parse_line: string -> ('a * string ATP_Problem.atp_problem_line list) list ->
   342 fun parse_class x = scan_general_id x
   342 fun parse_class x = scan_general_id x
   343 and parse_classes x = (parse_class ::: Scan.repeat ($$ "&" |-- parse_class)) x
   343 and parse_classes x = (parse_class ::: Scan.repeat ($$ "&" |-- parse_class)) x
   344 
   344 
   345 fun parse_type x =
   345 fun parse_type x =
   346   (($$ "(" |-- parse_type --| $$ ")"
   346   (($$ "(" |-- parse_type --| $$ ")"
       
   347     || Scan.this_string tptp_pi_binder |-- $$ "[" |-- skip_term --| $$ "]" --| $$ ":" -- parse_type
       
   348        >> (fn (_, ty) => ty (* currently ignoring type constructor declarations anyway *))
   347     || (scan_general_id -- Scan.optional ($$ "{" |-- parse_classes --| $$ "}") [])
   349     || (scan_general_id -- Scan.optional ($$ "{" |-- parse_classes --| $$ "}") [])
   348         -- Scan.optional ($$ "(" |-- parse_types --| $$ ")") []
   350         -- Scan.optional ($$ "(" |-- parse_types --| $$ ")") []
   349        >> AType)
   351        >> AType)
   350    -- Scan.option (($$ tptp_fun_type || $$ tptp_product_type) -- parse_type)
   352    -- Scan.option (($$ tptp_app || $$ tptp_fun_type || $$ tptp_product_type) -- parse_type)
   351   >> (fn (a, NONE) => a
   353   >> (fn (a, NONE) => a
   352        | (a, SOME (fun_or_product, b)) =>
   354        | (a, SOME (bin_op, b)) =>
   353          if fun_or_product = tptp_fun_type then AFun (a, b)
   355          if bin_op = tptp_app then
   354          else AType ((tptp_product_type, []), [a, b]))) x
   356            (case a of
       
   357              AType (s_clss, tys) => AType (s_clss, tys @ [b])
       
   358            | _ => raise UNRECOGNIZED_ATP_PROOF ())
       
   359          else if bin_op = tptp_fun_type then
       
   360            AFun (a, b)
       
   361          else if bin_op = tptp_product_type then
       
   362            AType ((tptp_product_type, []), [a, b])
       
   363          else
       
   364            raise Fail "impossible case")) x
   355 and parse_types x =
   365 and parse_types x =
   356   (parse_type ::: Scan.repeat ($$ "," |-- parse_type)) x
   366   (parse_type ::: Scan.repeat ($$ "," |-- parse_type)) x
   357 
   367 
   358 (* We currently half ignore types. *)
   368 (* We currently half ignore types. *)
   359 fun parse_optional_type_signature x =
   369 fun parse_optional_type_signature x =
   525 fun parse_typed_var x =
   535 fun parse_typed_var x =
   526   (Scan.repeat (scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type)
   536   (Scan.repeat (scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type)
   527      --| Scan.option (Scan.this_string ","))
   537      --| Scan.option (Scan.this_string ","))
   528    || $$ "(" |-- parse_typed_var --| $$ ")") x
   538    || $$ "(" |-- parse_typed_var --| $$ ")") x
   529 
   539 
   530 fun parse_simple_thf0_term x =
   540 fun parse_simple_thf_term x =
   531   (parse_fo_quantifier -- ($$ "[" |-- parse_typed_var --| $$ "]" --| $$ ":") -- parse_thf0_term
   541   (parse_fo_quantifier -- ($$ "[" |-- parse_typed_var --| $$ "]" --| $$ ":") -- parse_thf_term
   532       >> (fn ((q, ys), t) =>
   542       >> (fn ((q, ys), t) =>
   533           fold_rev
   543           fold_rev
   534             (fn (var, ty) => fn r =>
   544             (fn (var, ty) => fn r =>
   535                 AAbs (((var, the_default dummy_atype ty), r), [])
   545                 AAbs (((var, the_default dummy_atype ty), r), [])
   536                 |> (if tptp_lambda <> q then
   546                 |> (if tptp_lambda <> q then
   537                       mk_app (q |> mk_ho_of_fo_quant
   547                       mk_app (q |> mk_ho_of_fo_quant
   538                                 |> mk_simple_aterm)
   548                                 |> mk_simple_aterm)
   539                     else I))
   549                     else I))
   540             ys t)
   550             ys t)
   541   || Scan.this_string tptp_not |-- parse_thf0_term >> mk_app (mk_simple_aterm tptp_not)
   551   || Scan.this_string tptp_not |-- parse_thf_term >> mk_app (mk_simple_aterm tptp_not)
   542   || scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type)
   552   || scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type)
   543     >> (fn (var, typ_opt) => ATerm ((var, the_list typ_opt), []))
   553     >> (fn (var, typ_opt) => ATerm ((var, the_list typ_opt), []))
   544   || parse_ho_quantifier >> mk_simple_aterm
   554   || parse_ho_quantifier >> mk_simple_aterm
   545   || $$ "(" |-- parse_thf0_term --| $$ ")"
   555   || $$ "(" |-- parse_thf_term --| $$ ")"
   546   || parse_binary_op >> mk_simple_aterm) x
   556   || parse_binary_op >> mk_simple_aterm) x
   547 and parse_thf0_term x =
   557 and parse_thf_term x =
   548   (parse_simple_thf0_term -- Scan.option (parse_binary_op -- parse_thf0_term)
   558   (parse_simple_thf_term -- Scan.option (parse_binary_op -- parse_thf_term)
   549     >> (fn (t1, SOME (c, t2)) =>
   559     >> (fn (t1, SOME (c, t2)) =>
   550            if c = tptp_app then mk_app t1 t2 else mk_apps (mk_simple_aterm c) [t1, t2]
   560            if c = tptp_app then mk_app t1 t2 else mk_apps (mk_simple_aterm c) [t1, t2]
   551          | (t, NONE) => t)) x
   561          | (t, NONE) => t)) x
   552 
   562 
   553 fun parse_thf0_formula x = (parse_thf0_term #>> remove_thf_app #>> AAtom) x
   563 fun parse_thf_formula x = (parse_thf_term #>> remove_thf_app #>> AAtom) x
   554 
   564 
   555 fun parse_tstp_thf0_line problem =
   565 fun parse_tstp_thf_line problem =
   556   (Scan.this_string tptp_thf -- $$ "(") |-- scan_general_id --| $$ ","
   566   (Scan.this_string tptp_thf -- $$ "(") |-- scan_general_id --| $$ ","
   557   -- Symbol.scan_ascii_id --| $$ "," -- parse_thf0_formula -- parse_tstp_extra_arguments --| $$ ")"
   567   -- Symbol.scan_ascii_id --| $$ "," -- (parse_thf_formula || skip_term >> K dummy_phi)
       
   568      -- parse_tstp_extra_arguments --| $$ ")"
   558   --| $$ "."
   569   --| $$ "."
   559   >> (fn (((num, role), phi), deps) =>
   570   >> (fn (((num, role), phi), deps) =>
   560       let
   571       let
   561         val role' = role_of_tptp_string role
   572         val role' = role_of_tptp_string role
   562         val ((name, phi), rule, deps) =
   573         val ((name, phi), rule, deps) =
   681 fun parse_z3_tptp_core_line x =
   692 fun parse_z3_tptp_core_line x =
   682   (Scan.this_string "SZS core" |-- Scan.repeat ($$ " " |-- scan_general_id)
   693   (Scan.this_string "SZS core" |-- Scan.repeat ($$ " " |-- scan_general_id)
   683    >> map (core_inference z3_tptp_core_rule)) x
   694    >> map (core_inference z3_tptp_core_rule)) x
   684 
   695 
   685 fun parse_line local_name problem =
   696 fun parse_line local_name problem =
   686   if local_name = leo2N orelse local_name = leo3N then parse_tstp_thf0_line problem
   697   (* Satallax is handled separately, in "atp_satallax.ML". *)
       
   698   if local_name = leo2N orelse local_name = leo3N then parse_tstp_thf_line problem
   687   else if local_name = spassN then parse_spass_line
   699   else if local_name = spassN then parse_spass_line
   688   else if local_name = pirateN then parse_pirate_line
   700   else if local_name = pirateN then parse_pirate_line
   689   else if local_name = z3_tptpN then parse_z3_tptp_core_line
   701   else if local_name = z3_tptpN then parse_z3_tptp_core_line
   690   else parse_tstp_line problem
   702   else parse_tstp_line problem
   691 
   703