src/FOLP/FOLP.thy
author Simon Wimmer <wimmers@in.tum.de>
Thu, 18 Apr 2024 17:53:14 +0200
changeset 80137 0c51e0a6bc37
parent 69605 a96320074298
permissions -rw-r--r--
sketch & explore: recover from duplicate fixed variables in Isar proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1477
4c51ab632cda expanded tabs
clasohm
parents: 1142
diff changeset
     1
(*  Title:      FOLP/FOLP.thy
4c51ab632cda expanded tabs
clasohm
parents: 1142
diff changeset
     2
    Author:     Martin D Coen, Cambridge University Computer Laboratory
1142
eb0e2ff8f032 Corrected comments in headers
lcp
parents: 0
diff changeset
     3
    Copyright   1992  University of Cambridge
eb0e2ff8f032 Corrected comments in headers
lcp
parents: 0
diff changeset
     4
*)
eb0e2ff8f032 Corrected comments in headers
lcp
parents: 0
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
     6
section \<open>Classical First-Order Logic with Proofs\<close>
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
     7
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
     8
theory FOLP
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
     9
imports IFOLP
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
    10
begin
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
    11
41779
a68f503805ed modernized specifications;
wenzelm
parents: 36319
diff changeset
    12
axiomatization cla :: "[p=>p]=>p"
a68f503805ed modernized specifications;
wenzelm
parents: 36319
diff changeset
    13
  where classical: "(!!x. x:~P ==> f(x):P) ==> cla(f):P"
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
    14
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    15
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    16
(*** Classical introduction rules for | and EX ***)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    17
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    18
schematic_goal disjCI:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    19
  assumes "!!x. x:~Q ==> f(x):P"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    20
  shows "?p : P|Q"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    21
  apply (rule classical)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    22
  apply (assumption | rule assms disjI1 notI)+
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    23
  apply (assumption | rule disjI2 notE)+
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    24
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    25
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    26
(*introduction rule involving only EX*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    27
schematic_goal ex_classical:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    28
  assumes "!!u. u:~(EX x. P(x)) ==> f(u):P(a)"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    29
  shows "?p : EX x. P(x)"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    30
  apply (rule classical)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    31
  apply (rule exI, rule assms, assumption)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    32
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    33
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    34
(*version of above, simplifying ~EX to ALL~ *)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    35
schematic_goal exCI:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    36
  assumes "!!u. u:ALL x. ~P(x) ==> f(u):P(a)"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    37
  shows "?p : EX x. P(x)"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    38
  apply (rule ex_classical)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    39
  apply (rule notI [THEN allI, THEN assms])
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    40
  apply (erule notE)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    41
  apply (erule exI)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    42
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    43
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    44
schematic_goal excluded_middle: "?p : ~P | P"
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    45
  apply (rule disjCI)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    46
  apply assumption
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    47
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    48
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    49
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    50
(*** Special elimination rules *)
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
    51
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    52
(*Classical implies (-->) elimination. *)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    53
schematic_goal impCE:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    54
  assumes major: "p:P-->Q"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    55
    and r1: "!!x. x:~P ==> f(x):R"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    56
    and r2: "!!y. y:Q ==> g(y):R"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    57
  shows "?p : R"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    58
  apply (rule excluded_middle [THEN disjE])
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
    59
   apply (tactic \<open>DEPTH_SOLVE (assume_tac \<^context> 1 ORELSE
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
    60
       resolve_tac \<^context> [@{thm r1}, @{thm r2}, @{thm major} RS @{thm mp}] 1)\<close>)
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    61
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    62
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    63
(*Double negation law*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    64
schematic_goal notnotD: "p:~~P ==> ?p : P"
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    65
  apply (rule classical)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    66
  apply (erule notE)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    67
  apply assumption
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    68
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    69
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    70
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    71
(*** Tactics for implication and contradiction ***)
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
    72
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    73
(*Classical <-> elimination.  Proof substitutes P=Q in
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    74
    ~P ==> ~Q    and    P ==> Q  *)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    75
schematic_goal iffCE:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    76
  assumes major: "p:P<->Q"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    77
    and r1: "!!x y.[| x:P; y:Q |] ==> f(x,y):R"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    78
    and r2: "!!x y.[| x:~P; y:~Q |] ==> g(x,y):R"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    79
  shows "?p : R"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    80
  apply (insert major)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    81
  apply (unfold iff_def)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    82
  apply (rule conjE)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
    83
  apply (tactic \<open>DEPTH_SOLVE_1 (eresolve_tac \<^context> @{thms impCE} 1 ORELSE
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
    84
      eresolve_tac \<^context> [@{thm notE}, @{thm impE}] 1 THEN assume_tac \<^context> 1 ORELSE
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
    85
      assume_tac \<^context> 1 ORELSE
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
    86
      resolve_tac \<^context> [@{thm r1}, @{thm r2}] 1)\<close>)+
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    87
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    88
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    89
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    90
(*Should be used as swap since ~P becomes redundant*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    91
schematic_goal swap:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    92
  assumes major: "p:~P"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    93
    and r: "!!x. x:~Q ==> f(x):P"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    94
  shows "?p : Q"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    95
  apply (rule classical)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    96
  apply (rule major [THEN notE])
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    97
  apply (rule r)
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    98
  apply assumption
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
    99
  done
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   100
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   101
ML_file \<open>classical.ML\<close>      (* Patched because matching won't instantiate proof *)
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   102
ML_file \<open>simp.ML\<close>           (* Patched because matching won't instantiate proof *)
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   103
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   104
ML \<open>
42799
4e33894aec6d modernized functor names;
wenzelm
parents: 41779
diff changeset
   105
structure Cla = Classical
4e33894aec6d modernized functor names;
wenzelm
parents: 41779
diff changeset
   106
(
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   107
  val sizef = size_of_thm
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   108
  val mp = @{thm mp}
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   109
  val not_elim = @{thm notE}
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   110
  val swap = @{thm swap}
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   111
  fun hyp_subst_tacs ctxt = [hyp_subst_tac ctxt]
42799
4e33894aec6d modernized functor names;
wenzelm
parents: 41779
diff changeset
   112
);
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   113
open Cla;
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   114
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   115
(*Propositional rules
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   116
  -- iffCE might seem better, but in the examples in ex/cla
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   117
     run about 7% slower than with iffE*)
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   118
val prop_cs =
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   119
  empty_cs addSIs [@{thm refl}, @{thm TrueI}, @{thm conjI}, @{thm disjCI},
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   120
      @{thm impI}, @{thm notI}, @{thm iffI}]
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   121
    addSEs [@{thm conjE}, @{thm disjE}, @{thm impCE}, @{thm FalseE}, @{thm iffE}];
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   122
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   123
(*Quantifier rules*)
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   124
val FOLP_cs =
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   125
  prop_cs addSIs [@{thm allI}] addIs [@{thm exI}, @{thm ex1I}]
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   126
    addSEs [@{thm exE}, @{thm ex1E}] addEs [@{thm allE}];
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   127
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   128
val FOLP_dup_cs =
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   129
  prop_cs addSIs [@{thm allI}] addIs [@{thm exCI}, @{thm ex1I}]
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   130
    addSEs [@{thm exE}, @{thm ex1E}] addEs [@{thm all_dupE}];
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   131
\<close>
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   132
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   133
schematic_goal cla_rews:
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   134
  "?p1 : P | ~P"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   135
  "?p2 : ~P | P"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   136
  "?p3 : ~ ~ P <-> P"
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   137
  "?p4 : (~P --> P) <-> P"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61337
diff changeset
   138
  apply (tactic \<open>ALLGOALS (Cla.fast_tac \<^context> FOLP_cs)\<close>)
26322
eaf634e975fa converted legacy ML scripts;
wenzelm
parents: 22577
diff changeset
   139
  done
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   140
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   141
ML_file \<open>simpdata.ML\<close>
17480
fd19f77dcf60 converted to Isar theory format;
wenzelm
parents: 3836
diff changeset
   142
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
end