| author | wenzelm | 
| Sat, 15 Oct 2016 13:07:54 +0200 | |
| changeset 64221 | 407f69c4959f | 
| parent 61389 | 509d7ee638f8 | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 22809 | 1 | (* Title: Cube/Cube.thy | 
| 2 | Author: Tobias Nipkow | |
| 3 | *) | |
| 3773 | 4 | |
| 58889 | 5 | section \<open>Barendregt's Lambda-Cube\<close> | 
| 17252 | 6 | |
| 7 | theory Cube | |
| 8 | imports Pure | |
| 9 | begin | |
| 10 | ||
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
35256diff
changeset | 11 | setup Pure_Thy.old_appl_syntax_setup | 
| 26956 
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
 wenzelm parents: 
24783diff
changeset | 12 | |
| 57945 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
52143diff
changeset | 13 | named_theorems rules "Cube inference rules" | 
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 14 | |
| 17252 | 15 | typedecl "term" | 
| 16 | typedecl "context" | |
| 17 | typedecl typing | |
| 18 | ||
| 45241 | 19 | axiomatization | 
| 61387 | 20 | Abs :: "[term, term \<Rightarrow> term] \<Rightarrow> term" and | 
| 21 | Prod :: "[term, term \<Rightarrow> term] \<Rightarrow> term" and | |
| 22 | Trueprop :: "[context, typing] \<Rightarrow> prop" and | |
| 45241 | 23 | MT_context :: "context" and | 
| 61387 | 24 | Context :: "[typing, context] \<Rightarrow> context" and | 
| 45241 | 25 |   star :: "term"  ("*") and
 | 
| 45242 | 26 |   box :: "term"  ("\<box>") and
 | 
| 61388 | 27 | app :: "[term, term] \<Rightarrow> term" (infixl "\<cdot>" 20) and | 
| 61387 | 28 | Has_type :: "[term, term] \<Rightarrow> typing" | 
| 17252 | 29 | |
| 45241 | 30 | nonterminal context' and typing' | 
| 61387 | 31 | syntax | 
| 32 |   "_Trueprop" :: "[context', typing'] \<Rightarrow> prop"  ("(_/ \<turnstile> _)")
 | |
| 33 |   "_Trueprop1" :: "typing' \<Rightarrow> prop"  ("(_)")
 | |
| 34 |   "" :: "id \<Rightarrow> context'"  ("_")
 | |
| 35 |   "" :: "var \<Rightarrow> context'"  ("_")
 | |
| 36 |   "_MT_context" :: "context'"  ("")
 | |
| 37 |   "_Context" :: "[typing', context'] \<Rightarrow> context'"  ("_ _")
 | |
| 38 |   "_Has_type" :: "[term, term] \<Rightarrow> typing'"  ("(_:/ _)" [0, 0] 5)
 | |
| 61388 | 39 |   "_Lam" :: "[idt, term, term] \<Rightarrow> term"  ("(3\<^bold>\<lambda>_:_./ _)" [0, 0, 0] 10)
 | 
| 40 |   "_Pi" :: "[idt, term, term] \<Rightarrow> term"  ("(3\<Prod>_:_./ _)" [0, 0] 10)
 | |
| 61387 | 41 | "_arrow" :: "[term, term] \<Rightarrow> term" (infixr "\<rightarrow>" 10) | 
| 42 | translations | |
| 43 | "_Trueprop(G, t)" \<rightleftharpoons> "CONST Trueprop(G, t)" | |
| 44 |   ("prop") "x:X" \<rightleftharpoons> ("prop") "\<turnstile> x:X"
 | |
| 45 | "_MT_context" \<rightleftharpoons> "CONST MT_context" | |
| 46 | "_Context" \<rightleftharpoons> "CONST Context" | |
| 47 | "_Has_type" \<rightleftharpoons> "CONST Has_type" | |
| 61388 | 48 | "\<^bold>\<lambda>x:A. B" \<rightleftharpoons> "CONST Abs(A, \<lambda>x. B)" | 
| 49 | "\<Prod>x:A. B" \<rightharpoonup> "CONST Prod(A, \<lambda>x. B)" | |
| 61387 | 50 | "A \<rightarrow> B" \<rightharpoonup> "CONST Prod(A, \<lambda>_. B)" | 
| 58617 | 51 | print_translation \<open> | 
| 42284 | 52 |   [(@{const_syntax Prod},
 | 
| 52143 | 53 |     fn _ => Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"}))]
 | 
| 58617 | 54 | \<close> | 
| 17252 | 55 | |
| 45241 | 56 | axiomatization where | 
| 45242 | 57 | s_b: "*: \<box>" and | 
| 17252 | 58 | |
| 61387 | 59 | strip_s: "\<lbrakk>A:*; a:A \<Longrightarrow> G \<turnstile> x:X\<rbrakk> \<Longrightarrow> a:A G \<turnstile> x:X" and | 
| 60 | strip_b: "\<lbrakk>A:\<box>; a:A \<Longrightarrow> G \<turnstile> x:X\<rbrakk> \<Longrightarrow> a:A G \<turnstile> x:X" and | |
| 17252 | 61 | |
| 61388 | 62 | app: "\<lbrakk>F:Prod(A, B); C:A\<rbrakk> \<Longrightarrow> F\<cdot>C: B(C)" and | 
| 0 | 63 | |
| 61387 | 64 | pi_ss: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> B(x):*\<rbrakk> \<Longrightarrow> Prod(A, B):*" and | 
| 17252 | 65 | |
| 61387 | 66 | lam_ss: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):* \<rbrakk> | 
| 67 | \<Longrightarrow> Abs(A, f) : Prod(A, B)" and | |
| 17252 | 68 | |
| 61388 | 69 | beta: "Abs(A, f)\<cdot>a \<equiv> f(a)" | 
| 17252 | 70 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 71 | lemmas [rules] = s_b strip_s strip_b app lam_ss pi_ss | 
| 17252 | 72 | |
| 73 | lemma imp_elim: | |
| 61388 | 74 | assumes "f:A\<rightarrow>B" and "a:A" and "f\<cdot>a:B \<Longrightarrow> PROP P" | 
| 41526 | 75 | shows "PROP P" by (rule app assms)+ | 
| 17252 | 76 | |
| 77 | lemma pi_elim: | |
| 61388 | 78 | assumes "F:Prod(A,B)" and "a:A" and "F\<cdot>a:B(a) \<Longrightarrow> PROP P" | 
| 41526 | 79 | shows "PROP P" by (rule app assms)+ | 
| 17252 | 80 | |
| 81 | ||
| 82 | locale L2 = | |
| 61387 | 83 | assumes pi_bs: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> B(x):*\<rbrakk> \<Longrightarrow> Prod(A,B):*" | 
| 84 | and lam_bs: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):*\<rbrakk> | |
| 85 | \<Longrightarrow> Abs(A,f) : Prod(A,B)" | |
| 45242 | 86 | begin | 
| 17252 | 87 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 88 | lemmas [rules] = lam_bs pi_bs | 
| 45242 | 89 | |
| 90 | end | |
| 17252 | 91 | |
| 17260 | 92 | |
| 17252 | 93 | locale Lomega = | 
| 94 | assumes | |
| 61387 | 95 | pi_bb: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> \<Longrightarrow> Prod(A,B):\<box>" | 
| 96 | and lam_bb: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> | |
| 97 | \<Longrightarrow> Abs(A,f) : Prod(A,B)" | |
| 45242 | 98 | begin | 
| 17252 | 99 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 100 | lemmas [rules] = lam_bb pi_bb | 
| 45242 | 101 | |
| 102 | end | |
| 17252 | 103 | |
| 104 | ||
| 105 | locale LP = | |
| 61387 | 106 | assumes pi_sb: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> \<Longrightarrow> Prod(A,B):\<box>" | 
| 107 | and lam_sb: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> | |
| 108 | \<Longrightarrow> Abs(A,f) : Prod(A,B)" | |
| 45241 | 109 | begin | 
| 17252 | 110 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 111 | lemmas [rules] = lam_sb pi_sb | 
| 45241 | 112 | |
| 113 | end | |
| 17252 | 114 | |
| 17260 | 115 | |
| 17252 | 116 | locale LP2 = LP + L2 | 
| 45241 | 117 | begin | 
| 17252 | 118 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 119 | lemmas [rules] = lam_bs pi_bs lam_sb pi_sb | 
| 45241 | 120 | |
| 121 | end | |
| 17252 | 122 | |
| 17260 | 123 | |
| 17252 | 124 | locale Lomega2 = L2 + Lomega | 
| 45241 | 125 | begin | 
| 17252 | 126 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 127 | lemmas [rules] = lam_bs pi_bs lam_bb pi_bb | 
| 45241 | 128 | |
| 129 | end | |
| 17252 | 130 | |
| 17260 | 131 | |
| 17252 | 132 | locale LPomega = LP + Lomega | 
| 45241 | 133 | begin | 
| 17252 | 134 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 135 | lemmas [rules] = lam_bb pi_bb lam_sb pi_sb | 
| 45241 | 136 | |
| 137 | end | |
| 17252 | 138 | |
| 17260 | 139 | |
| 17252 | 140 | locale CC = L2 + LP + Lomega | 
| 45241 | 141 | begin | 
| 17252 | 142 | |
| 49752 
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
 wenzelm parents: 
45242diff
changeset | 143 | lemmas [rules] = lam_bs pi_bs lam_bb pi_bb lam_sb pi_sb | 
| 17252 | 144 | |
| 145 | end | |
| 45241 | 146 | |
| 147 | end |