src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML
author wenzelm
Wed, 08 Oct 2014 10:15:04 +0200
changeset 58629 a6a6cd499d4e
parent 55508 90c42b130652
child 59580 cbc38731d42f
permissions -rw-r--r--
tuned signature; tuned whitespace;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41474
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 37744
diff changeset
     1
(*  Title:      HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 37744
diff changeset
     2
    Author:     Philipp Meyer, TU Muenchen
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
     3
41474
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 37744
diff changeset
     4
Functions for generating a certificate from a positivstellensatz and vice
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 37744
diff changeset
     5
versa.
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
     6
*)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
     7
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
     8
signature POSITIVSTELLENSATZ_TOOLS =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
     9
sig
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    10
  val print_cert: RealArith.pss_tree -> string
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    11
  val read_cert: Proof.context -> string -> RealArith.pss_tree
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    12
end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    13
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    14
structure Positivstellensatz_Tools : POSITIVSTELLENSATZ_TOOLS =
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    15
struct
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    16
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    17
(** print certificate **)
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    18
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    19
local
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    20
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    21
fun string_of_rat r =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    22
  let
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    23
    val (nom, den) = Rat.quotient_of_rat r
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    24
  in
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    25
    if den = 1 then string_of_int nom
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    26
    else string_of_int nom ^ "/" ^ string_of_int den
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    27
  end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    28
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    29
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    30
(* map polynomials to strings *)
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    31
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    32
fun string_of_varpow x k =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    33
  let
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    34
    val term = term_of x
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    35
    val name =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    36
      (case term of
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    37
        Free (n, _) => n
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    38
      | _ => error "Term in monomial not free variable")
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    39
  in
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    40
    if k = 1 then name else name ^ "^" ^ string_of_int k
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    41
  end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    42
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    43
fun string_of_monomial m =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    44
  if FuncUtil.Ctermfunc.is_empty m then "1"
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    45
  else
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    46
    let
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    47
      val m' = FuncUtil.dest_monomial m
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    48
      val vps = fold_rev (fn (x,k) => cons (string_of_varpow x k)) m' []
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    49
    in foldr1 (fn (s, t) => s ^ "*" ^ t) vps end
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    50
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    51
fun string_of_cmonomial (m,c) =
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
    52
  if FuncUtil.Ctermfunc.is_empty m then string_of_rat c
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    53
  else if c = Rat.one then string_of_monomial m
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    54
  else string_of_rat c ^ "*" ^ string_of_monomial m
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    55
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    56
fun string_of_poly p =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    57
  if FuncUtil.Monomialfunc.is_empty p then "0"
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    58
  else
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    59
    let
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    60
      val cms = map string_of_cmonomial
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    61
        (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    62
    in foldr1 (fn (t1, t2) => t1 ^ " + " ^ t2) cms end
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    63
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    64
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    65
(* print cert *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    66
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    67
fun pss_to_cert (RealArith.Axiom_eq i) = "A=" ^ string_of_int i
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    68
  | pss_to_cert (RealArith.Axiom_le i) = "A<=" ^ string_of_int i
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    69
  | pss_to_cert (RealArith.Axiom_lt i) = "A<" ^ string_of_int i
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    70
  | pss_to_cert (RealArith.Rational_eq r) = "R=" ^ string_of_rat r
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    71
  | pss_to_cert (RealArith.Rational_le r) = "R<=" ^ string_of_rat r
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    72
  | pss_to_cert (RealArith.Rational_lt r) = "R<" ^ string_of_rat r
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
    73
  | pss_to_cert (RealArith.Square p) = "[" ^ string_of_poly p ^ "]^2"
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    74
  | pss_to_cert (RealArith.Eqmul (p, pss)) =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    75
      "([" ^ string_of_poly p ^ "] * " ^ pss_to_cert pss ^ ")"
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    76
  | pss_to_cert (RealArith.Sum (pss1, pss2)) =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    77
      "(" ^ pss_to_cert pss1 ^ " + " ^ pss_to_cert pss2 ^ ")"
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    78
  | pss_to_cert (RealArith.Product (pss1, pss2)) =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    79
      "(" ^ pss_to_cert pss1 ^ " * " ^ pss_to_cert pss2 ^ ")"
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    80
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    81
in
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    82
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    83
fun print_cert RealArith.Trivial = "()"
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    84
  | print_cert (RealArith.Cert pss) = "(" ^ pss_to_cert pss ^ ")"
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    85
  | print_cert (RealArith.Branch (t1, t2)) =
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    86
      "(" ^ print_cert t1 ^ " & " ^ print_cert t2 ^ ")"
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    87
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    88
end
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
    89
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    90
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    91
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    92
(** read certificate **)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    93
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    94
local
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    95
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
    96
(* basic parsers *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    97
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
    98
val str = Scan.this_string
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
    99
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   100
val number =
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   101
  Scan.repeat1 (Scan.one Symbol.is_ascii_digit >> (fn s => ord s - ord "0"))
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   102
    >> foldl1 (fn (n, d) => n * 10 + d)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   103
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   104
val nat = number
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   105
val int = Scan.optional (str "~" >> K ~1) 1 -- nat >> op *
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   106
val rat = int --| str "/" -- int >> Rat.rat_of_quotient
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   107
val rat_int = rat || int >> Rat.rat_of_int
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   108
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   109
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   110
(* polynomial parsers *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   111
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   112
fun repeat_sep s f = f ::: Scan.repeat (str s |-- f)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   113
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   114
val parse_id = Scan.one Symbol.is_letter ::: Scan.many Symbol.is_letdig >> implode
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   115
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   116
fun parse_varpow ctxt = parse_id -- Scan.optional (str "^" |-- nat) 1 >>
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   117
  (fn (x, k) => (cterm_of (Proof_Context.theory_of ctxt) (Free (x, @{typ real})), k))
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   118
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   119
fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >>
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33030
diff changeset
   120
  (fn xs => fold FuncUtil.Ctermfunc.update xs FuncUtil.Ctermfunc.empty)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   121
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   122
fun parse_cmonomial ctxt =
32646
962b4354ed90 used standard fold function and type aliases
Philipp Meyer
parents: 32645
diff changeset
   123
  rat_int --| str "*" -- (parse_monomial ctxt) >> swap ||
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   124
  (parse_monomial ctxt) >> (fn m => (m, Rat.one)) ||
32829
671eb46eb0a3 tuned FuncFun and FuncUtil structure in positivstellensatz.ML
Philipp Meyer
parents: 32828
diff changeset
   125
  rat_int >> (fn r => (FuncUtil.Ctermfunc.empty, r))
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   126
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   127
fun parse_poly ctxt = repeat_sep "+" (parse_cmonomial ctxt) >>
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33030
diff changeset
   128
  (fn xs => fold FuncUtil.Monomialfunc.update xs FuncUtil.Monomialfunc.empty)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   129
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   130
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   131
(* positivstellensatz parsers *)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   132
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   133
val parse_axiom =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
   134
  (str "A=" |-- int >> RealArith.Axiom_eq) ||
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
   135
  (str "A<=" |-- int >> RealArith.Axiom_le) ||
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
   136
  (str "A<" |-- int >> RealArith.Axiom_lt)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   137
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   138
val parse_rational =
32828
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
   139
  (str "R=" |-- rat_int >> RealArith.Rational_eq) ||
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
   140
  (str "R<=" |-- rat_int >> RealArith.Rational_le) ||
ad76967c703d removed opening of structures
Philipp Meyer
parents: 32646
diff changeset
   141
  (str "R<" |-- rat_int >> RealArith.Rational_lt)
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   142
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   143
fun parse_cert ctxt input =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   144
  let
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   145
    val pc = parse_cert ctxt
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   146
    val pp = parse_poly ctxt
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   147
  in
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   148
    (parse_axiom ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   149
     parse_rational ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   150
     str "[" |-- pp --| str "]^2" >> RealArith.Square ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   151
     str "([" |-- pp --| str "]*" -- pc --| str ")" >> RealArith.Eqmul ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   152
     str "(" |-- pc --| str "*" -- pc --| str ")" >> RealArith.Product ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   153
     str "(" |-- pc --| str "+" -- pc --| str ")" >> RealArith.Sum) input
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   154
  end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   155
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   156
fun parse_cert_tree ctxt input =
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   157
  let
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   158
    val pc = parse_cert ctxt
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   159
    val pt = parse_cert_tree ctxt
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   160
  in
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   161
    (str "()" >> K RealArith.Trivial ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   162
     str "(" |-- pc --| str ")" >> RealArith.Cert ||
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   163
     str "(" |-- pt --| str "&" -- pt --| str ")" >> RealArith.Branch) input
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   164
  end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   165
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   166
in
55508
90c42b130652 tuned whitespace;
wenzelm
parents: 43946
diff changeset
   167
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   168
fun read_cert ctxt input_str =
43946
wenzelm
parents: 42361
diff changeset
   169
  Symbol.scanner "Bad certificate" (parse_cert_tree ctxt)
wenzelm
parents: 42361
diff changeset
   170
    (filter_out Symbol.is_blank (Symbol.explode input_str))
32645
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   171
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   172
end
1cc5b24f5a01 sos method generates and uses proof certificates
Philipp Meyer
parents:
diff changeset
   173
58629
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   174
end
a6a6cd499d4e tuned signature;
wenzelm
parents: 55508
diff changeset
   175