src/HOL/ex/Iff_Oracle.thy
author haftmann
Tue, 13 Oct 2015 09:21:15 +0200
changeset 61424 c3658c18b7bc
parent 60754 02924903a6fd
child 69597 ff784d5a5bfb
permissions -rw-r--r--
prod_case as canonical name for product type eliminator
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
3f51f0945a3e Example of declaring oracles
paulson
parents:
diff changeset
     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
3f51f0945a3e Example of declaring oracles
paulson
parents:
diff changeset
     4
*)
3f51f0945a3e Example of declaring oracles
paulson
parents:
diff changeset
     5
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
     6
section \<open>Example of Declaring an Oracle\<close>
16063
7dd4eb2c8055 oracle example converted to Isar
paulson
parents: 3817
diff changeset
     7
29752
ad4e3a577fd3 modernized some theory names;
wenzelm
parents: 28316
diff changeset
     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
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    10
begin
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    11
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    12
subsection \<open>Oracle declaration\<close>
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    13
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    14
text \<open>
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    15
  This oracle makes tautologies of the form @{prop "P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P"}.
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    16
  The length is specified by an integer, which is checked to be even
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    17
  and positive.
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    18
\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    19
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    20
oracle iff_oracle = \<open>
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    21
  let
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
    22
    fun mk_iff 1 = Var (("P", 0), @{typ bool})
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
    23
      | mk_iff n = HOLogic.mk_eq (Var (("P", 0), @{typ bool}), mk_iff (n - 1));
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    24
  in
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 27809
diff changeset
    25
    fn (thy, n) =>
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    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
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    28
      else raise Fail ("iff_oracle: " ^ string_of_int n)
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    29
  end
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    30
\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    31
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    32
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    33
subsection \<open>Oracle as low-level rule\<close>
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    34
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    35
ML \<open>iff_oracle (@{theory}, 2)\<close>
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    36
ML \<open>iff_oracle (@{theory}, 10)\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    37
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    38
ML \<open>
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    39
  Thm.peek_status (iff_oracle (@{theory}, 10));
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    40
  @{assert} (#oracle it);
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    41
\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    42
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    43
text \<open>These oracle calls had better fail.\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    44
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    45
ML \<open>
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    46
  (iff_oracle (@{theory}, 5); error "Bad oracle")
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    47
    handle Fail _ => writeln "Oracle failed, as expected"
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    48
\<close>
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    49
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    50
ML \<open>
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
    51
  (iff_oracle (@{theory}, 1); error "Bad oracle")
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    52
    handle Fail _ => writeln "Oracle failed, as expected"
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    53
\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    54
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    55
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    56
subsection \<open>Oracle as proof method\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    57
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    58
method_setup iff =
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    59
  \<open>Scan.lift Parse.nat >> (fn n => fn ctxt =>
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 29752
diff changeset
    60
    SIMPLE_METHOD
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59621
diff changeset
    61
      (HEADGOAL (resolve_tac ctxt [iff_oracle (Proof_Context.theory_of ctxt, n)])
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    62
        handle Fail _ => no_tac))\<close>
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    63
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    64
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    65
lemma "A \<longleftrightarrow> A"
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    66
  by (iff 2)
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    67
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    68
lemma "A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A"
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    69
  by (iff 10)
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    70
59030
67baff6f697c misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    71
lemma "A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A \<longleftrightarrow> A"
16832
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    72
  apply (iff 5)?
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    73
  oops
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    74
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    75
lemma A
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    76
  apply (iff 1)?
f220d1df0f4e new type-safe interface;
wenzelm
parents: 16417
diff changeset
    77
  oops
3817
f20f193d42b4 removed declIffOracle;
wenzelm
parents: 1537
diff changeset
    78
16063
7dd4eb2c8055 oracle example converted to Isar
paulson
parents: 3817
diff changeset
    79
end