| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 70565 | d0b75c59beca | 
| child 71465 | 910a081cca74 | 
| permissions | -rw-r--r-- | 
| 
40239
 
c4336e45f199
moved FOL/ex/Iff_Oracle.thy to HOL/ex where it is more accessible to most readers of isar-ref;
 
wenzelm 
parents: 
36960 
diff
changeset
 | 
1  | 
(* Title: HOL/ex/Iff_Oracle.thy  | 
| 1537 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
40239
 
c4336e45f199
moved FOL/ex/Iff_Oracle.thy to HOL/ex where it is more accessible to most readers of isar-ref;
 
wenzelm 
parents: 
36960 
diff
changeset
 | 
3  | 
Author: Makarius  | 
| 1537 | 4  | 
*)  | 
5  | 
||
| 59030 | 6  | 
section \<open>Example of Declaring an Oracle\<close>  | 
| 16063 | 7  | 
|
| 29752 | 8  | 
theory Iff_Oracle  | 
| 
40239
 
c4336e45f199
moved FOL/ex/Iff_Oracle.thy to HOL/ex where it is more accessible to most readers of isar-ref;
 
wenzelm 
parents: 
36960 
diff
changeset
 | 
9  | 
imports Main  | 
| 16832 | 10  | 
begin  | 
| 3817 | 11  | 
|
| 59030 | 12  | 
subsection \<open>Oracle declaration\<close>  | 
| 16832 | 13  | 
|
| 59030 | 14  | 
text \<open>  | 
| 69597 | 15  | 
This oracle makes tautologies of the form \<^prop>\<open>P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P\<close>.  | 
| 16832 | 16  | 
The length is specified by an integer, which is checked to be even  | 
17  | 
and positive.  | 
|
| 59030 | 18  | 
\<close>  | 
| 3817 | 19  | 
|
| 59030 | 20  | 
oracle iff_oracle = \<open>  | 
| 16832 | 21  | 
let  | 
| 69597 | 22  | 
    fun mk_iff 1 = Var (("P", 0), \<^typ>\<open>bool\<close>)
 | 
23  | 
      | mk_iff n = HOLogic.mk_eq (Var (("P", 0), \<^typ>\<open>bool\<close>), mk_iff (n - 1));
 | 
|
| 16832 | 24  | 
in  | 
| 28290 | 25  | 
fn (thy, n) =>  | 
| 16832 | 26  | 
if n > 0 andalso n mod 2 = 0  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59030 
diff
changeset
 | 
27  | 
then Thm.global_cterm_of thy (HOLogic.mk_Trueprop (mk_iff n))  | 
| 16832 | 28  | 
      else raise Fail ("iff_oracle: " ^ string_of_int n)
 | 
29  | 
end  | 
|
| 59030 | 30  | 
\<close>  | 
| 3817 | 31  | 
|
| 16832 | 32  | 
|
| 59030 | 33  | 
subsection \<open>Oracle as low-level rule\<close>  | 
34  | 
||
| 69597 | 35  | 
ML \<open>iff_oracle (\<^theory>, 2)\<close>  | 
36  | 
ML \<open>iff_oracle (\<^theory>, 10)\<close>  | 
|
| 3817 | 37  | 
|
| 59030 | 38  | 
ML \<open>  | 
| 70565 | 39  | 
\<^assert> (map #1 (Thm_Deps.all_oracles [iff_oracle (\<^theory>, 10)]) = [\<^oracle_name>\<open>iff_oracle\<close>]);  | 
| 59030 | 40  | 
\<close>  | 
| 3817 | 41  | 
|
| 59030 | 42  | 
text \<open>These oracle calls had better fail.\<close>  | 
| 3817 | 43  | 
|
| 59030 | 44  | 
ML \<open>  | 
| 69597 | 45  | 
(iff_oracle (\<^theory>, 5); error "Bad oracle")  | 
| 59030 | 46  | 
handle Fail _ => writeln "Oracle failed, as expected"  | 
47  | 
\<close>  | 
|
48  | 
||
49  | 
ML \<open>  | 
|
| 69597 | 50  | 
(iff_oracle (\<^theory>, 1); error "Bad oracle")  | 
| 59030 | 51  | 
handle Fail _ => writeln "Oracle failed, as expected"  | 
52  | 
\<close>  | 
|
| 3817 | 53  | 
|
| 16832 | 54  | 
|
| 59030 | 55  | 
subsection \<open>Oracle as proof method\<close>  | 
| 3817 | 56  | 
|
| 59030 | 57  | 
method_setup iff =  | 
58  | 
\<open>Scan.lift Parse.nat >> (fn n => fn ctxt =>  | 
|
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
29752 
diff
changeset
 | 
59  | 
SIMPLE_METHOD  | 
| 60754 | 60  | 
(HEADGOAL (resolve_tac ctxt [iff_oracle (Proof_Context.theory_of ctxt, n)])  | 
| 59030 | 61  | 
handle Fail _ => no_tac))\<close>  | 
| 3817 | 62  | 
|
| 16832 | 63  | 
|
| 59030 | 64  | 
lemma "A \<longleftrightarrow> A"  | 
| 16832 | 65  | 
by (iff 2)  | 
| 3817 | 66  | 
|
| 59030 | 67  | 
lemma "A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A"  | 
| 16832 | 68  | 
by (iff 10)  | 
69  | 
||
| 59030 | 70  | 
lemma "A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A"  | 
| 16832 | 71  | 
apply (iff 5)?  | 
72  | 
oops  | 
|
73  | 
||
74  | 
lemma A  | 
|
75  | 
apply (iff 1)?  | 
|
76  | 
oops  | 
|
| 3817 | 77  | 
|
| 16063 | 78  | 
end  |