src/Sequents/Sequents.thy
author haftmann
Tue, 20 Jul 2010 06:35:29 +0200
changeset 37891 c26f9d06e82c
parent 35430 df2862dc23a8
child 38499 8f0cd11238a7
permissions -rw-r--r--
robustified metis proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     1
(*  Title:      Sequents/Sequents.thy
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     4
*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     5
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     6
header {* Parsing and pretty-printing of sequences *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     7
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     8
theory Sequents
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     9
imports Pure
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    10
uses ("prover.ML")
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    11
begin
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    12
26956
1309a6a0a29f setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents: 24178
diff changeset
    13
setup PureThy.old_appl_syntax_setup
1309a6a0a29f setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents: 24178
diff changeset
    14
24178
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 18176
diff changeset
    15
declare [[unify_trace_bound = 20, unify_search_bound = 40]]
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 18176
diff changeset
    16
7098
86583034aacf installation of simplifier and classical reasoner, better rules etc
paulson
parents: 6819
diff changeset
    17
global
86583034aacf installation of simplifier and classical reasoner, better rules etc
paulson
parents: 6819
diff changeset
    18
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    19
typedecl o
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    20
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    21
(* Sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    23
typedecl
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    24
 seq'
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    25
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    26
consts
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    27
 SeqO'         :: "[o,seq']=>seq'"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    28
 Seq1'         :: "o=>seq'"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    29
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    30
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    31
(* concrete syntax *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    32
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7166
diff changeset
    33
nonterminals
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    34
  seq seqobj seqcont
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    35
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    36
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    37
syntax
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    38
 "_SeqEmp"         :: seq                                  ("")
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    39
 "_SeqApp"         :: "[seqobj,seqcont] => seq"            ("__")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    40
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    41
 "_SeqContEmp"     :: seqcont                              ("")
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    42
 "_SeqContApp"     :: "[seqobj,seqcont] => seqcont"        (",/ __")
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    43
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    44
 "_SeqO"           :: "o => seqobj"                        ("_")
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    45
 "_SeqId"          :: "'a => seqobj"                       ("$_")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    46
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    47
types
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    48
 single_seqe = "[seq,seqobj] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    49
 single_seqi = "[seq'=>seq',seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    50
 two_seqi    = "[seq'=>seq', seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    51
 two_seqe    = "[seq, seq] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    52
 three_seqi  = "[seq'=>seq', seq'=>seq', seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    53
 three_seqe  = "[seq, seq, seq] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    54
 four_seqi   = "[seq'=>seq', seq'=>seq', seq'=>seq', seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    55
 four_seqe   = "[seq, seq, seq, seq] => prop"
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    56
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    57
 sequence_name = "seq'=>seq'"
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    58
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    59
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7166
diff changeset
    60
syntax
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    61
  (*Constant to allow definitions of SEQUENCES of formulas*)
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    62
  "_Side"        :: "seq=>(seq'=>seq')"     ("<<(_)>>")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    63
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    64
ML {*
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    65
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    66
(* parse translation for sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    67
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35363
diff changeset
    68
fun abs_seq' t = Abs ("s", Type (@{type_name seq'}, []), t);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    69
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    70
fun seqobj_tr (Const (@{syntax_const "_SeqO"}, _) $ f) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    71
      Const (@{const_syntax SeqO'}, dummyT) $ f
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    72
  | seqobj_tr (_ $ i) = i;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    73
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    74
fun seqcont_tr (Const (@{syntax_const "_SeqContEmp"}, _)) = Bound 0
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    75
  | seqcont_tr (Const (@{syntax_const "_SeqContApp"}, _) $ so $ sc) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    76
      seqobj_tr so $ seqcont_tr sc;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    77
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    78
fun seq_tr (Const (@{syntax_const "_SeqEmp"}, _)) = abs_seq' (Bound 0)
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    79
  | seq_tr (Const (@{syntax_const "_SeqApp"}, _) $ so $ sc) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    80
      abs_seq'(seqobj_tr so $ seqcont_tr sc);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    81
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    82
fun singlobj_tr (Const (@{syntax_const "_SeqO"},_) $ f) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    83
  abs_seq' ((Const (@{const_syntax SeqO'}, dummyT) $ f) $ Bound 0);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    84
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    85
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    86
(* print translation for sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    87
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    88
fun seqcont_tr' (Bound 0) =
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    89
      Const (@{syntax_const "_SeqContEmp"}, dummyT)
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    90
  | seqcont_tr' (Const (@{const_syntax SeqO'}, _) $ f $ s) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    91
      Const (@{syntax_const "_SeqContApp"}, dummyT) $
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    92
        (Const (@{syntax_const "_SeqO"}, dummyT) $ f) $ seqcont_tr' s
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    93
  | seqcont_tr' (i $ s) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    94
      Const (@{syntax_const "_SeqContApp"}, dummyT) $
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    95
        (Const (@{syntax_const "_SeqId"}, dummyT) $ i) $ seqcont_tr' s;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    96
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    97
fun seq_tr' s =
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    98
  let
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
    99
    fun seq_itr' (Bound 0) = Const (@{syntax_const "_SeqEmp"}, dummyT)
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   100
      | seq_itr' (Const (@{const_syntax SeqO'}, _) $ f $ s) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   101
          Const (@{syntax_const "_SeqApp"}, dummyT) $
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   102
            (Const (@{syntax_const "_SeqO"}, dummyT) $ f) $ seqcont_tr' s
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   103
      | seq_itr' (i $ s) =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   104
          Const (@{syntax_const "_SeqApp"}, dummyT) $
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   105
            (Const (@{syntax_const "_SeqId"}, dummyT) $ i) $ seqcont_tr' s
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   106
  in
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   107
    case s of
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   108
      Abs (_, _, t) => seq_itr' t
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   109
    | _ => s $ Bound 0
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   110
  end;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   111
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   112
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   113
fun single_tr c [s1, s2] =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   114
  Const (c, dummyT) $ seq_tr s1 $ singlobj_tr s2;
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   115
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   116
fun two_seq_tr c [s1, s2] =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   117
  Const (c, dummyT) $ seq_tr s1 $ seq_tr s2;
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   118
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   119
fun three_seq_tr c [s1, s2, s3] =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   120
  Const (c, dummyT) $ seq_tr s1 $ seq_tr s2 $ seq_tr s3;
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   121
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   122
fun four_seq_tr c [s1, s2, s3, s4] =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   123
  Const (c, dummyT) $ seq_tr s1 $ seq_tr s2 $ seq_tr s3 $ seq_tr s4;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   124
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   125
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   126
fun singlobj_tr' (Const (@{const_syntax SeqO'},_) $ fm) = fm
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   127
  | singlobj_tr' id = Const (@{syntax_const "_SeqId"}, dummyT) $ id;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   128
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   129
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   130
fun single_tr' c [s1, s2] =
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   131
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   132
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   133
fun two_seq_tr' c [s1, s2] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   134
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   135
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   136
fun three_seq_tr' c [s1, s2, s3] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   137
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   138
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   139
fun four_seq_tr' c [s1, s2, s3, s4] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   140
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3 $ seq_tr' s4;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   141
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   142
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   143
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   144
(** for the <<...>> notation **)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   145
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   146
fun side_tr [s1] = seq_tr s1;
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   147
*}
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   148
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 26956
diff changeset
   149
parse_translation {* [(@{syntax_const "_Side"}, side_tr)] *}
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   150
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   151
use "prover.ML"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   152
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   153
end
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   154