src/HOL/HOL.thy
author oheimb
Tue, 23 Apr 1996 16:58:21 +0200
changeset 1672 2c109cd2fdd0
parent 1370 7361ac9b024d
child 1674 33aff4d854e4
permissions -rw-r--r--
repaired critical proofs depending on the order inside non-confluent SimpSets, (temporarily) removed problematic rule less_Suc_eq form simpset_of "Nat"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/HOL.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Higher-Order Logic
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
HOL = CPure +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
classes
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
  term < logic
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
axclass
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
  plus < term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
axclass
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
  minus < term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
axclass
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
  times < term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
default
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
  term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
types
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
  bool
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
arities
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
  fun :: (term, term) term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
  bool :: term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
  (* Constants *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    38
  Trueprop      :: bool => prop                     ("(_)" 5)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    39
  not           :: bool => bool                     ("~ _" [40] 40)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    40
  True, False   :: bool
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    41
  If            :: [bool, 'a, 'a] => 'a   ("(if (_)/ then (_)/ else (_))" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    42
  Inv           :: ('a => 'b) => ('b => 'a)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
  (* Binders *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    46
  Eps           :: ('a => bool) => 'a
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    47
  All           :: ('a => bool) => bool             (binder "! " 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    48
  Ex            :: ('a => bool) => bool             (binder "? " 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    49
  Ex1           :: ('a => bool) => bool             (binder "?! " 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    50
  Let           :: ['a, 'a => 'b] => 'b
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
  (* Infixes *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    54
  o             :: ['b => 'c, 'a => 'b, 'a] => 'c   (infixl 55)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    55
  "="           :: ['a, 'a] => bool                 (infixl 50)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    56
  "&"           :: [bool, bool] => bool             (infixr 35)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    57
  "|"           :: [bool, bool] => bool             (infixr 30)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    58
  "-->"         :: [bool, bool] => bool             (infixr 25)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
  (* Overloaded Constants *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    62
  "+"           :: ['a::plus, 'a] => 'a             (infixl 65)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    63
  "-"           :: ['a::minus, 'a] => 'a            (infixl 65)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    64
  "*"           :: ['a::times, 'a] => 'a            (infixl 70)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
types
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
  letbinds  letbind
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
  case_syn  cases_syn
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
syntax
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    73
  "~="          :: ['a, 'a] => bool                 (infixl 50)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    75
  "@Eps"        :: [pttrn,bool] => 'a               ("(3@ _./ _)" 10)
1068
e0f2dffab506 HOL.thy:
nipkow
parents: 973
diff changeset
    76
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
  (* Alternative Quantifiers *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    79
  "*All"        :: [idts, bool] => bool             ("(3ALL _./ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    80
  "*Ex"         :: [idts, bool] => bool             ("(3EX _./ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    81
  "*Ex1"        :: [idts, bool] => bool             ("(3EX! _./ _)" 10)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
  (* Let expressions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    85
  "_bind"       :: [pttrn, 'a] => letbind           ("(2_ =/ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    86
  ""            :: letbind => letbinds              ("_")
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    87
  "_binds"      :: [letbind, letbinds] => letbinds  ("_;/ _")
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    88
  "_Let"        :: [letbinds, 'a] => 'a             ("(let (_)/ in (_))" 10)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
  (* Case expressions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    92
  "@case"       :: ['a, cases_syn] => 'b            ("(case _ of/ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    93
  "@case1"      :: ['a, 'b] => case_syn             ("(2_ =>/ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    94
  ""            :: case_syn => cases_syn            ("_")
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    95
  "@case2"      :: [case_syn, cases_syn] => cases_syn   ("_/ | _")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
translations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
  "x ~= y"      == "~ (x = y)"
1068
e0f2dffab506 HOL.thy:
nipkow
parents: 973
diff changeset
    99
  "@ x.b"       == "Eps(%x.b)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
  "ALL xs. P"   => "! xs. P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
  "EX xs. P"    => "? xs. P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
  "EX! xs. P"   => "?! xs. P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
  "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
1114
c8dfb56a7e95 Prod is now a parent of Lfp.
nipkow
parents: 1068
diff changeset
   104
  "let x = a in e"        == "Let a (%x. e)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
rules
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
  eq_reflection "(x=y) ==> (x==y)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
  (* Basic Rules *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
  refl          "t = (t::'a)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
  subst         "[| s = t; P(s) |] ==> P(t::'a)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
  ext           "(!!x::'a. (f(x)::'b) = g(x)) ==> (%x.f(x)) = (%x.g(x))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
  selectI       "P(x::'a) ==> P(@x.P(x))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
  impI          "(P ==> Q) ==> P-->Q"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
  mp            "[| P-->Q;  P |] ==> Q"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
  True_def      "True      == ((%x::bool.x)=(%x.x))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
  All_def       "All(P)    == (P = (%x.True))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
  Ex_def        "Ex(P)     == P(@x.P(x))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
  False_def     "False     == (!P.P)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
  not_def       "~ P       == P-->False"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
  and_def       "P & Q     == !R. (P-->Q-->R) --> R"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
  or_def        "P | Q     == !R. (P-->R) --> (Q-->R) --> R"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
  Ex1_def       "Ex1(P)    == ? x. P(x) & (! y. P(y) --> y=x)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
rules
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
  (* Axioms *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
  iff           "(P-->Q) --> (Q-->P) --> (P=Q)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
  True_or_False "(P=True) | (P=False)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
  (* Misc Definitions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
  Let_def       "Let s f == f(s)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
  Inv_def       "Inv(f::'a=>'b)  == (% y. @x. f(x)=y)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
  o_def         "(f::'b=>'c) o g == (%(x::'a). f(g(x)))"
973
f57fb576520f Modified If_def to avoid ambiguity.
nipkow
parents: 965
diff changeset
   143
  if_def        "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
1273
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   145
(* start 8bit 1 *)
1672
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   146
types
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   147
('a, 'b) "ë"          (infixr 5)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   148
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   149
syntax
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   150
  "Ú"		:: "['a::{}, 'a] => prop"       ("(_ Ú/ _)"         [3, 2] 2)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   151
  "êë"		:: "[prop, prop] => prop"	("(_ êë/ _)"        [2, 1] 1)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   152
  "Gbigimpl"	:: "[asms, prop] => prop"	("((3Ë _ Ì) êë/ _)" [0, 1] 1)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   153
  "Gall"	:: "('a => prop) => prop"	(binder "Ä"                0)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   154
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   155
  "Glam"         :: "[idts, 'a::logic] => 'b::logic"  ("(3³_./ _)" 10)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   156
  "Û"           :: "['a, 'a] => bool"                 (infixl 50)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   157
  "Gnot"        :: "bool => bool"                     ("¿ _" [40] 40)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   158
  "GEps"        :: "[pttrn, bool] => 'a"               ("(3®_./ _)" 10)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   159
  "GAll"        :: "[idts, bool] => bool"             ("(3Â_./ _)" 10)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   160
  "GEx"         :: "[idts, bool] => bool"             ("(3Ã_./ _)" 10)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   161
  "GEx1"        :: "[idts, bool] => bool"             ("(3Ã!_./ _)" 10)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   162
  "À"           :: "[bool, bool] => bool"             (infixr 35)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   163
  "Á"           :: "[bool, bool] => bool"             (infixr 30)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   164
  "çè"          :: "[bool, bool] => bool"             (infixr 25)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   165
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   166
translations
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   167
(type)  "x ë y"	== (type) "x => y" 
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   168
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   169
(*  "³x.t"	=> "%x.t" *)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   170
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   171
  "x Û y"	== "x ~= y"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   172
  "¿ y"		== "~y"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   173
  "®x.P"	== "@x.P"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   174
  "Âx.P"	== "! x. P"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   175
  "Ãx.P"	== "? x. P"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   176
  "Ã!x.P"	== "?! x. P"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   177
  "x À y"	== "x & y"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   178
  "x Á y"	== "x | y"
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   179
  "x çè y"	== "x --> y"
1273
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   180
(* end 8bit 1 *)
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   181
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
end
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
ML
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
(** Choice between the HOL and Isabelle style of quantifiers **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
val HOL_quantifiers = ref true;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
fun alt_ast_tr' (name, alt_name) =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
  let
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
    fun ast_tr' (*name*) args =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
      if ! HOL_quantifiers then raise Match
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
      else Syntax.mk_appl (Syntax.Constant alt_name) args;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
  in
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
    (name, ast_tr')
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
  end;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
val print_ast_translation =
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
  map alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")];
1273
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   202
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   203
(* start 8bit 2 *)
1672
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   204
local open Ast;
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   205
fun bigimpl_ast_tr (*"Gbigimpl"*) [asms, concl] =
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   206
      fold_ast_p "êë" (unfold_ast "_asms" asms, concl)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   207
  | bigimpl_ast_tr (*"Gbigimpl"*) asts = raise_ast "bigimpl_ast_tr" asts;
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   208
fun impl_ast_tr' (*"êë"*) asts =
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   209
  (case unfold_ast_p "êë" (Appl (Constant "êë" :: asts)) of
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   210
    (asms as _ :: _ :: _, concl)
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   211
      => Appl [Constant "Gbigimpl", fold_ast "_asms" asms, concl]
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   212
  | _ => raise Match);
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   213
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   214
in
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   215
val parse_ast_translation = ("Gbigimpl", bigimpl_ast_tr)::
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   216
				("Glam",    fn asts => Appl (Constant "_abs" :: asts))::
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   217
				parse_ast_translation;
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   218
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   219
val print_ast_translation = ("êë", impl_ast_tr')::
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   220
				("_lambda", fn asts => Appl (Constant "Glam" :: asts)) ::	
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   221
				print_ast_translation;
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   222
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   223
end;
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   224
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   225
local open Syntax in
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   226
val thy = thy 
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   227
|> add_trrules_i 
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   228
[mk_appl (Constant "Ú" ) [Variable "P", Variable "Q"] <-> 
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   229
 mk_appl (Constant "==") [Variable "P", Variable "Q"],
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   230
 mk_appl (Constant "êë" ) [Variable "P", Variable "Q"] <-> 
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   231
 mk_appl (Constant "==>") [Variable "P", Variable "Q"],
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   232
 (Constant "Ä" ) <->  (Constant "!!")]
2c109cd2fdd0 repaired critical proofs depending on the order inside non-confluent SimpSets,
oheimb
parents: 1370
diff changeset
   233
end;
1273
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   234
(* end 8bit 2 *)
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   235
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   236
6960ec882bca added 8bit pragmas
regensbu
parents: 1232
diff changeset
   237