changeset 70830 | 8f050cc0ec50 |
parent 70449 | 6e34025981be |
child 70839 | 2136e4670ad2 |
--- a/src/HOL/Proofs/ex/Proof_Terms.thy Fri Oct 11 11:16:36 2019 +0200 +++ b/src/HOL/Proofs/ex/Proof_Terms.thy Fri Oct 11 15:36:32 2019 +0200 @@ -23,7 +23,7 @@ val thm = @{thm ex}; (*proof body with digest*) - val body = Proofterm.strip_thm (Thm.proof_body_of thm); + val body = Proofterm.strip_thm_body (Thm.proof_body_of thm); (*proof term only*) val prf = Proofterm.proof_of body;