src/FOL/fologic.ML
author paulson <lp15@cam.ac.uk>
Wed, 24 Apr 2024 20:56:26 +0100
changeset 80149 40a3fc07a587
parent 74342 5d411d85da8c
permissions -rw-r--r--
More tidying of proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     1
(*  Title:      FOL/fologic.ML
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     3
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     4
Abstract syntax operations for FOL.
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     5
*)
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     6
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     7
signature FOLOGIC =
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
     8
sig
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
     9
  val mk_conj: term * term -> term
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    10
  val mk_disj: term * term -> term
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    11
  val mk_imp: term * term -> term
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    12
  val dest_imp: term -> term * term
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    13
  val dest_conj: term -> term list
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    14
  val mk_iff: term * term -> term
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    15
  val dest_iff: term -> term * term
74321
714e87ce6e9d clarified signature;
wenzelm
parents: 74320
diff changeset
    16
  val mk_all: string * typ -> term -> term
714e87ce6e9d clarified signature;
wenzelm
parents: 74320
diff changeset
    17
  val mk_exists: string * typ -> term -> term
32449
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    18
  val mk_eq: term * term -> term
696d64ed85da eliminated hard tabs;
wenzelm
parents: 31974
diff changeset
    19
  val dest_eq: term -> term * term
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    20
end;
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    21
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    22
structure FOLogic: FOLOGIC =
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    23
struct
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    24
74319
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    25
fun mk_conj (t1, t2) = \<^Const>\<open>conj for t1 t2\<close>
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    26
and mk_disj (t1, t2) = \<^Const>\<open>disj for t1 t2\<close>
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    27
and mk_imp (t1, t2) = \<^Const>\<open>imp for t1 t2\<close>
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    28
and mk_iff (t1, t2) = \<^Const>\<open>iff for t1 t2\<close>;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    29
74319
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    30
val dest_imp = \<^Const_fn>\<open>imp for A B => \<open>(A, B)\<close>\<close>;
4466
305390f23734 Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents: 4349
diff changeset
    31
74293
54279cfcf037 more antiquotations;
wenzelm
parents: 69593
diff changeset
    32
fun dest_conj \<^Const_>\<open>conj for t t'\<close> = t :: dest_conj t'
11668
548ba68385a3 added dest_conj, dest_concls;
wenzelm
parents: 10384
diff changeset
    33
  | dest_conj t = [t];
548ba68385a3 added dest_conj, dest_concls;
wenzelm
parents: 10384
diff changeset
    34
74319
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    35
val dest_iff = \<^Const_fn>\<open>iff for A B => \<open>(A, B)\<close>\<close>;
9543
ce61d1c1a509 new abstract syntax operations, used in ZF
paulson
parents: 9473
diff changeset
    36
74320
dd04da556d1a clarified antiquotations;
wenzelm
parents: 74319
diff changeset
    37
fun mk_eq (t, u) =
dd04da556d1a clarified antiquotations;
wenzelm
parents: 74319
diff changeset
    38
  let val T = fastype_of t
dd04da556d1a clarified antiquotations;
wenzelm
parents: 74319
diff changeset
    39
  in \<^Const>\<open>eq T for t u\<close> end;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    40
74319
54b2e5f771da clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents: 74293
diff changeset
    41
val dest_eq = \<^Const_fn>\<open>eq _ for lhs rhs => \<open>(lhs, rhs)\<close>\<close>;
6140
af32e2c3f77a tidied; added dest_eq
paulson
parents: 4466
diff changeset
    42
74321
714e87ce6e9d clarified signature;
wenzelm
parents: 74320
diff changeset
    43
fun mk_all (x, T) P = \<^Const>\<open>All T for \<open>absfree (x, T) P\<close>\<close>;
714e87ce6e9d clarified signature;
wenzelm
parents: 74320
diff changeset
    44
fun mk_exists (x, T) P = \<^Const>\<open>Ex T for \<open>absfree (x, T) P\<close>\<close>;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    45
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
diff changeset
    46
end;