src/Cube/Cube.thy
author wenzelm
Sun, 30 Nov 2008 14:43:29 +0100
changeset 28917 20f43e0e0958
parent 26956 1309a6a0a29f
child 35054 a5db9779b026
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22809
3cf5df73d50a added header;
wenzelm
parents: 17782
diff changeset
     1
(*  Title:      Cube/Cube.thy
3cf5df73d50a added header;
wenzelm
parents: 17782
diff changeset
     2
    ID:         $Id$
3cf5df73d50a added header;
wenzelm
parents: 17782
diff changeset
     3
    Author:     Tobias Nipkow
3cf5df73d50a added header;
wenzelm
parents: 17782
diff changeset
     4
*)
3773
989ef5e9d543 syntactic constants;
wenzelm
parents: 1149
diff changeset
     5
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
     6
header {* Barendregt's Lambda-Cube *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
     7
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
     8
theory Cube
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
     9
imports Pure
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    10
begin
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    11
26956
1309a6a0a29f setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents: 24783
diff changeset
    12
setup PureThy.old_appl_syntax_setup
1309a6a0a29f setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents: 24783
diff changeset
    13
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    14
typedecl "term"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    15
typedecl "context"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    16
typedecl typing
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    17
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    18
nonterminals
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    19
  context' typing'
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    20
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    21
consts
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    22
  Abs           :: "[term, term => term] => term"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    23
  Prod          :: "[term, term => term] => term"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    24
  Trueprop      :: "[context, typing] => prop"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    25
  MT_context    :: "context"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    26
  Context       :: "[typing, context] => context"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    27
  star          :: "term"                               ("*")
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    28
  box           :: "term"                               ("[]")
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    29
  app           :: "[term, term] => term"               (infixl "^" 20)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    30
  Has_type      :: "[term, term] => typing"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    31
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    32
syntax
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    33
  Trueprop      :: "[context', typing'] => prop"        ("(_/ |- _)")
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    34
  Trueprop1     :: "typing' => prop"                    ("(_)")
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    35
  ""            :: "id => context'"                     ("_")
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    36
  ""            :: "var => context'"                    ("_")
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    37
  MT_context    :: "context'"                           ("")
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    38
  Context       :: "[typing', context'] => context'"    ("_ _")
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    39
  Has_type      :: "[term, term] => typing'"            ("(_:/ _)" [0, 0] 5)
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    40
  Lam           :: "[idt, term, term] => term"          ("(3Lam _:_./ _)" [0, 0, 0] 10)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    41
  Pi            :: "[idt, term, term] => term"          ("(3Pi _:_./ _)" [0, 0] 10)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    42
  arrow         :: "[term, term] => term"               (infixr "->" 10)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    43
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    44
translations
17260
wenzelm
parents: 17252
diff changeset
    45
  ("prop") "x:X" == ("prop") "|- x:X"
wenzelm
parents: 17252
diff changeset
    46
  "Lam x:A. B"   == "Abs(A, %x. B)"
wenzelm
parents: 17252
diff changeset
    47
  "Pi x:A. B"    => "Prod(A, %x. B)"
17782
b3846df9d643 replaced _K by dummy abstraction;
wenzelm
parents: 17260
diff changeset
    48
  "A -> B"       => "Prod(A, %_. B)"
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    49
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    50
syntax (xsymbols)
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 22809
diff changeset
    51
  Trueprop      :: "[context', typing'] => prop"        ("(_/ \<turnstile> _)")
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    52
  box           :: "term"                               ("\<box>")
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    53
  Lam           :: "[idt, term, term] => term"          ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    54
  Pi            :: "[idt, term, term] => term"          ("(3\<Pi> _:_./ _)" [0, 0] 10)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    55
  arrow         :: "[term, term] => term"               (infixr "\<rightarrow>" 10)
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    56
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    57
print_translation {* [("Prod", dependent_tr' ("Pi", "arrow"))] *}
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    58
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    59
axioms
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    60
  s_b:          "*: []"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    61
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    62
  strip_s:      "[| A:*;  a:A ==> G |- x:X |] ==> a:A G |- x:X"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    63
  strip_b:      "[| A:[]; a:A ==> G |- x:X |] ==> a:A G |- x:X"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    64
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    65
  app:          "[| F:Prod(A, B); C:A |] ==> F^C: B(C)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    67
  pi_ss:        "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    68
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    69
  lam_ss:       "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    70
                   ==> Abs(A, f) : Prod(A, B)"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    71
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    72
  beta:          "Abs(A, f)^a == f(a)"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    73
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    74
lemmas simple = s_b strip_s strip_b app lam_ss pi_ss
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    75
lemmas rules = simple
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    76
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    77
lemma imp_elim:
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    78
  assumes "f:A->B" and "a:A" and "f^a:B ==> PROP P"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    79
  shows "PROP P" by (rule app prems)+
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    80
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    81
lemma pi_elim:
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    82
  assumes "F:Prod(A,B)" and "a:A" and "F^a:B(a) ==> PROP P"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    83
  shows "PROP P" by (rule app prems)+
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    84
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    85
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    86
locale L2 =
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    87
  assumes pi_bs: "[| A:[]; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    88
    and lam_bs: "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    89
                   ==> Abs(A,f) : Prod(A,B)"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    90
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    91
lemmas (in L2) rules = simple lam_bs pi_bs
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    92
17260
wenzelm
parents: 17252
diff changeset
    93
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    94
locale Lomega =
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    95
  assumes
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    96
    pi_bb: "[| A:[]; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    97
    and lam_bb: "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |]
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    98
                   ==> Abs(A,f) : Prod(A,B)"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
    99
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   100
lemmas (in Lomega) rules = simple lam_bb pi_bb
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   101
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   102
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   103
locale LP =
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   104
  assumes pi_sb: "[| A:*; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   105
    and lam_sb: "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |]
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   106
                   ==> Abs(A,f) : Prod(A,B)"
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   107
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   108
lemmas (in LP) rules = simple lam_sb pi_sb
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   109
17260
wenzelm
parents: 17252
diff changeset
   110
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   111
locale LP2 = LP + L2
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   112
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   113
lemmas (in LP2) rules = simple lam_bs pi_bs lam_sb pi_sb
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   114
17260
wenzelm
parents: 17252
diff changeset
   115
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   116
locale Lomega2 = L2 + Lomega
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   117
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   118
lemmas (in Lomega2) rules = simple lam_bs pi_bs lam_bb pi_bb
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   119
17260
wenzelm
parents: 17252
diff changeset
   120
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   121
locale LPomega = LP + Lomega
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   122
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   123
lemmas (in LPomega) rules = simple lam_bb pi_bb lam_sb pi_sb
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   124
17260
wenzelm
parents: 17252
diff changeset
   125
17252
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   126
locale CC = L2 + LP + Lomega
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   127
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   128
lemmas (in CC) rules = simple lam_bs pi_bs lam_bb pi_bb lam_sb pi_sb
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   129
e352f65d5893 converted to Isar theory format;
wenzelm
parents: 11260
diff changeset
   130
end