src/HOL/Library/Eval_Witness.thy
changeset 27368 9f90ac19e32b
parent 27103 d8549f4d900b
child 27487 c8a6ce181805
equal deleted inserted replaced
27367:a75d71c73362 27368:9f90ac19e32b
     5 *)
     5 *)
     6 
     6 
     7 header {* Evaluation Oracle with ML witnesses *}
     7 header {* Evaluation Oracle with ML witnesses *}
     8 
     8 
     9 theory Eval_Witness
     9 theory Eval_Witness
    10 imports List
    10 imports Plain List
    11 begin
    11 begin
    12 
    12 
    13 text {* 
    13 text {* 
    14   We provide an oracle method similar to "eval", but with the
    14   We provide an oracle method similar to "eval", but with the
    15   possibility to provide ML values as witnesses for existential
    15   possibility to provide ML values as witnesses for existential