src/Cube/Cube.thy
changeset 45242 401f91ed8a93
parent 45241 87950f752099
child 49752 2bbb0013ff82
equal deleted inserted replaced
45241:87950f752099 45242:401f91ed8a93
    19   Prod :: "[term, term => term] => term" and
    19   Prod :: "[term, term => term] => term" and
    20   Trueprop :: "[context, typing] => prop" and
    20   Trueprop :: "[context, typing] => prop" and
    21   MT_context :: "context" and
    21   MT_context :: "context" and
    22   Context :: "[typing, context] => context" and
    22   Context :: "[typing, context] => context" and
    23   star :: "term"  ("*") and
    23   star :: "term"  ("*") and
    24   box :: "term"  ("[]") and
    24   box :: "term"  ("\<box>") and
    25   app :: "[term, term] => term"  (infixl "^" 20) and
    25   app :: "[term, term] => term"  (infixl "^" 20) and
    26   Has_type :: "[term, term] => typing"
    26   Has_type :: "[term, term] => typing"
    27 
       
    28 notation (xsymbols)
       
    29   box  ("\<box>")
       
    30 
    27 
    31 nonterminal context' and typing'
    28 nonterminal context' and typing'
    32 
    29 
    33 syntax
    30 syntax
    34   "_Trueprop" :: "[context', typing'] => prop"  ("(_/ |- _)")
    31   "_Trueprop" :: "[context', typing'] => prop"  ("(_/ \<turnstile> _)")
    35   "_Trueprop1" :: "typing' => prop"  ("(_)")
    32   "_Trueprop1" :: "typing' => prop"  ("(_)")
    36   "" :: "id => context'"  ("_")
    33   "" :: "id => context'"  ("_")
    37   "" :: "var => context'"  ("_")
    34   "" :: "var => context'"  ("_")
    38   "_MT_context" :: "context'"  ("")
    35   "_MT_context" :: "context'"  ("")
    39   "_Context" :: "[typing', context'] => context'"  ("_ _")
    36   "_Context" :: "[typing', context'] => context'"  ("_ _")
    40   "_Has_type" :: "[term, term] => typing'"  ("(_:/ _)" [0, 0] 5)
    37   "_Has_type" :: "[term, term] => typing'"  ("(_:/ _)" [0, 0] 5)
    41   "_Lam" :: "[idt, term, term] => term"  ("(3Lam _:_./ _)" [0, 0, 0] 10)
    38   "_Lam" :: "[idt, term, term] => term"  ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
    42   "_Pi" :: "[idt, term, term] => term"  ("(3Pi _:_./ _)" [0, 0] 10)
    39   "_Pi" :: "[idt, term, term] => term"  ("(3\<Pi> _:_./ _)" [0, 0] 10)
    43   "_arrow" :: "[term, term] => term"  (infixr "->" 10)
    40   "_arrow" :: "[term, term] => term"  (infixr "\<rightarrow>" 10)
    44 
    41 
    45 translations
    42 translations
    46   "_Trueprop(G, t)" == "CONST Trueprop(G, t)"
    43   "_Trueprop(G, t)" == "CONST Trueprop(G, t)"
    47   ("prop") "x:X" == ("prop") "|- x:X"
    44   ("prop") "x:X" == ("prop") "\<turnstile> x:X"
    48   "_MT_context" == "CONST MT_context"
    45   "_MT_context" == "CONST MT_context"
    49   "_Context" == "CONST Context"
    46   "_Context" == "CONST Context"
    50   "_Has_type" == "CONST Has_type"
    47   "_Has_type" == "CONST Has_type"
    51   "Lam x:A. B" == "CONST Abs(A, %x. B)"
    48   "\<Lambda> x:A. B" == "CONST Abs(A, %x. B)"
    52   "Pi x:A. B" => "CONST Prod(A, %x. B)"
    49   "\<Pi> x:A. B" => "CONST Prod(A, %x. B)"
    53   "A -> B" => "CONST Prod(A, %_. B)"
    50   "A \<rightarrow> B" => "CONST Prod(A, %_. B)"
    54 
    51 
    55 syntax (xsymbols)
    52 syntax (xsymbols)
    56   "_Trueprop" :: "[context', typing'] => prop"    ("(_/ \<turnstile> _)")
       
    57   "_Lam" :: "[idt, term, term] => term"    ("(3\<Lambda> _:_./ _)" [0, 0, 0] 10)
       
    58   "_Pi" :: "[idt, term, term] => term"    ("(3\<Pi> _:_./ _)" [0, 0] 10)
    53   "_Pi" :: "[idt, term, term] => term"    ("(3\<Pi> _:_./ _)" [0, 0] 10)
    59   "_arrow" :: "[term, term] => term"    (infixr "\<rightarrow>" 10)
       
    60 
    54 
    61 print_translation {*
    55 print_translation {*
    62   [(@{const_syntax Prod},
    56   [(@{const_syntax Prod},
    63     Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"}))]
    57     Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"}))]
    64 *}
    58 *}
    65 
    59 
    66 axiomatization where
    60 axiomatization where
    67   s_b: "*: []"  and
    61   s_b: "*: \<box>"  and
    68 
    62 
    69   strip_s: "[| A:*;  a:A ==> G |- x:X |] ==> a:A G |- x:X" and
    63   strip_s: "[| A:*;  a:A ==> G \<turnstile> x:X |] ==> a:A G \<turnstile> x:X" and
    70   strip_b: "[| A:[]; a:A ==> G |- x:X |] ==> a:A G |- x:X" and
    64   strip_b: "[| A:\<box>; a:A ==> G \<turnstile> x:X |] ==> a:A G \<turnstile> x:X" and
    71 
    65 
    72   app: "[| F:Prod(A, B); C:A |] ==> F^C: B(C)" and
    66   app: "[| F:Prod(A, B); C:A |] ==> F^C: B(C)" and
    73 
    67 
    74   pi_ss: "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*" and
    68   pi_ss: "[| A:*; !!x. x:A ==> B(x):* |] ==> Prod(A, B):*" and
    75 
    69 
    80 
    74 
    81 lemmas simple = s_b strip_s strip_b app lam_ss pi_ss
    75 lemmas simple = s_b strip_s strip_b app lam_ss pi_ss
    82 lemmas rules = simple
    76 lemmas rules = simple
    83 
    77 
    84 lemma imp_elim:
    78 lemma imp_elim:
    85   assumes "f:A->B" and "a:A" and "f^a:B ==> PROP P"
    79   assumes "f:A\<rightarrow>B" and "a:A" and "f^a:B ==> PROP P"
    86   shows "PROP P" by (rule app assms)+
    80   shows "PROP P" by (rule app assms)+
    87 
    81 
    88 lemma pi_elim:
    82 lemma pi_elim:
    89   assumes "F:Prod(A,B)" and "a:A" and "F^a:B(a) ==> PROP P"
    83   assumes "F:Prod(A,B)" and "a:A" and "F^a:B(a) ==> PROP P"
    90   shows "PROP P" by (rule app assms)+
    84   shows "PROP P" by (rule app assms)+
    91 
    85 
    92 
    86 
    93 locale L2 =
    87 locale L2 =
    94   assumes pi_bs: "[| A:[]; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"
    88   assumes pi_bs: "[| A:\<box>; !!x. x:A ==> B(x):* |] ==> Prod(A,B):*"
    95     and lam_bs: "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
    89     and lam_bs: "[| A:\<box>; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):* |]
    96                    ==> Abs(A,f) : Prod(A,B)"
    90                    ==> Abs(A,f) : Prod(A,B)"
       
    91 begin
    97 
    92 
    98 lemmas (in L2) rules = simple lam_bs pi_bs
    93 lemmas rules = simple lam_bs pi_bs
       
    94 
       
    95 end
    99 
    96 
   100 
    97 
   101 locale Lomega =
    98 locale Lomega =
   102   assumes
    99   assumes
   103     pi_bb: "[| A:[]; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"
   100     pi_bb: "[| A:\<box>; !!x. x:A ==> B(x):\<box> |] ==> Prod(A,B):\<box>"
   104     and lam_bb: "[| A:[]; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |]
   101     and lam_bb: "[| A:\<box>; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):\<box> |]
   105                    ==> Abs(A,f) : Prod(A,B)"
   102                    ==> Abs(A,f) : Prod(A,B)"
       
   103 begin
   106 
   104 
   107 lemmas (in Lomega) rules = simple lam_bb pi_bb
   105 lemmas rules = simple lam_bb pi_bb
       
   106 
       
   107 end
   108 
   108 
   109 
   109 
   110 locale LP =
   110 locale LP =
   111   assumes pi_sb: "[| A:*; !!x. x:A ==> B(x):[] |] ==> Prod(A,B):[]"
   111   assumes pi_sb: "[| A:*; !!x. x:A ==> B(x):\<box> |] ==> Prod(A,B):\<box>"
   112     and lam_sb: "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):[] |]
   112     and lam_sb: "[| A:*; !!x. x:A ==> f(x):B(x); !!x. x:A ==> B(x):\<box> |]
   113                    ==> Abs(A,f) : Prod(A,B)"
   113                    ==> Abs(A,f) : Prod(A,B)"
   114 begin
   114 begin
   115 
   115 
   116 lemmas rules = simple lam_sb pi_sb
   116 lemmas rules = simple lam_sb pi_sb
   117 
   117