author | haftmann |
Wed, 21 May 2025 20:13:43 +0200 | |
changeset 82648 | 35e40c60c680 |
parent 80923 | 6c9628a116cc |
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:
35256
diff
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:
24783
diff
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:
52143
diff
changeset
|
13 |
named_theorems rules "Cube inference rules" |
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
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 |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
25 |
star :: "term" (\<open>*\<close>) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
26 |
box :: "term" (\<open>\<box>\<close>) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
27 |
app :: "[term, term] \<Rightarrow> term" (infixl \<open>\<cdot>\<close> 20) and |
61387 | 28 |
Has_type :: "[term, term] \<Rightarrow> typing" |
17252 | 29 |
|
45241 | 30 |
nonterminal context' and typing' |
61387 | 31 |
syntax |
80923 | 32 |
"_Trueprop" :: "[context', typing'] \<Rightarrow> prop" (\<open>(\<open>notation=judgment\<close>_/ \<turnstile> _)\<close>) |
33 |
"_Trueprop1" :: "typing' \<Rightarrow> prop" (\<open>(\<open>notation=judgment\<close>_)\<close>) |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
34 |
"" :: "id \<Rightarrow> context'" (\<open>_\<close>) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
35 |
"" :: "var \<Rightarrow> context'" (\<open>_\<close>) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
36 |
"_MT_context" :: "context'" (\<open>\<close>) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
37 |
"_Context" :: "[typing', context'] \<Rightarrow> context'" (\<open>_ _\<close>) |
80917 | 38 |
"_Has_type" :: "[term, term] \<Rightarrow> typing'" (\<open>(\<open>notation=\<open>infix Has_Type\<close>\<close>_:/ _)\<close> [0, 0] 5) |
39 |
"_Lam" :: "[idt, term, term] \<Rightarrow> term" (\<open>(\<open>indent=3 notation=\<open>binder \<^bold>\<lambda>\<close>\<close>\<^bold>\<lambda>_:_./ _)\<close> [0, 0, 0] 10) |
|
40 |
"_Pi" :: "[idt, term, term] \<Rightarrow> term" (\<open>(\<open>indent=3 notation=\<open>binder \<Prod>\<close>\<close>\<Prod>_:_./ _)\<close> [0, 0] 10) |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80761
diff
changeset
|
41 |
"_arrow" :: "[term, term] \<Rightarrow> term" (infixr \<open>\<rightarrow>\<close> 10) |
80761 | 42 |
syntax_consts |
43 |
"_Trueprop" \<rightleftharpoons> Trueprop and |
|
44 |
"_MT_context" \<rightleftharpoons> MT_context and |
|
45 |
"_Context" \<rightleftharpoons> Context and |
|
46 |
"_Has_type" \<rightleftharpoons> Has_type and |
|
47 |
"_Lam" \<rightleftharpoons> Abs and |
|
48 |
"_Pi" "_arrow" \<rightleftharpoons> Prod |
|
61387 | 49 |
translations |
50 |
"_Trueprop(G, t)" \<rightleftharpoons> "CONST Trueprop(G, t)" |
|
51 |
("prop") "x:X" \<rightleftharpoons> ("prop") "\<turnstile> x:X" |
|
52 |
"_MT_context" \<rightleftharpoons> "CONST MT_context" |
|
53 |
"_Context" \<rightleftharpoons> "CONST Context" |
|
54 |
"_Has_type" \<rightleftharpoons> "CONST Has_type" |
|
61388 | 55 |
"\<^bold>\<lambda>x:A. B" \<rightleftharpoons> "CONST Abs(A, \<lambda>x. B)" |
56 |
"\<Prod>x:A. B" \<rightharpoonup> "CONST Prod(A, \<lambda>x. B)" |
|
61387 | 57 |
"A \<rightarrow> B" \<rightharpoonup> "CONST Prod(A, \<lambda>_. B)" |
58617 | 58 |
print_translation \<open> |
69593 | 59 |
[(\<^const_syntax>\<open>Prod\<close>, |
60 |
fn _ => Syntax_Trans.dependent_tr' (\<^syntax_const>\<open>_Pi\<close>, \<^syntax_const>\<open>_arrow\<close>))] |
|
58617 | 61 |
\<close> |
17252 | 62 |
|
45241 | 63 |
axiomatization where |
45242 | 64 |
s_b: "*: \<box>" and |
17252 | 65 |
|
61387 | 66 |
strip_s: "\<lbrakk>A:*; a:A \<Longrightarrow> G \<turnstile> x:X\<rbrakk> \<Longrightarrow> a:A G \<turnstile> x:X" and |
67 |
strip_b: "\<lbrakk>A:\<box>; a:A \<Longrightarrow> G \<turnstile> x:X\<rbrakk> \<Longrightarrow> a:A G \<turnstile> x:X" and |
|
17252 | 68 |
|
61388 | 69 |
app: "\<lbrakk>F:Prod(A, B); C:A\<rbrakk> \<Longrightarrow> F\<cdot>C: B(C)" and |
0 | 70 |
|
61387 | 71 |
pi_ss: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> B(x):*\<rbrakk> \<Longrightarrow> Prod(A, B):*" and |
17252 | 72 |
|
61387 | 73 |
lam_ss: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):* \<rbrakk> |
74 |
\<Longrightarrow> Abs(A, f) : Prod(A, B)" and |
|
17252 | 75 |
|
61388 | 76 |
beta: "Abs(A, f)\<cdot>a \<equiv> f(a)" |
17252 | 77 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
78 |
lemmas [rules] = s_b strip_s strip_b app lam_ss pi_ss |
17252 | 79 |
|
80 |
lemma imp_elim: |
|
61388 | 81 |
assumes "f:A\<rightarrow>B" and "a:A" and "f\<cdot>a:B \<Longrightarrow> PROP P" |
41526 | 82 |
shows "PROP P" by (rule app assms)+ |
17252 | 83 |
|
84 |
lemma pi_elim: |
|
61388 | 85 |
assumes "F:Prod(A,B)" and "a:A" and "F\<cdot>a:B(a) \<Longrightarrow> PROP P" |
41526 | 86 |
shows "PROP P" by (rule app assms)+ |
17252 | 87 |
|
88 |
||
89 |
locale L2 = |
|
61387 | 90 |
assumes pi_bs: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> B(x):*\<rbrakk> \<Longrightarrow> Prod(A,B):*" |
91 |
and lam_bs: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):*\<rbrakk> |
|
92 |
\<Longrightarrow> Abs(A,f) : Prod(A,B)" |
|
45242 | 93 |
begin |
17252 | 94 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
95 |
lemmas [rules] = lam_bs pi_bs |
45242 | 96 |
|
97 |
end |
|
17252 | 98 |
|
17260 | 99 |
|
17252 | 100 |
locale Lomega = |
101 |
assumes |
|
61387 | 102 |
pi_bb: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> \<Longrightarrow> Prod(A,B):\<box>" |
103 |
and lam_bb: "\<lbrakk>A:\<box>; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> |
|
104 |
\<Longrightarrow> Abs(A,f) : Prod(A,B)" |
|
45242 | 105 |
begin |
17252 | 106 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
107 |
lemmas [rules] = lam_bb pi_bb |
45242 | 108 |
|
109 |
end |
|
17252 | 110 |
|
111 |
||
112 |
locale LP = |
|
61387 | 113 |
assumes pi_sb: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> \<Longrightarrow> Prod(A,B):\<box>" |
114 |
and lam_sb: "\<lbrakk>A:*; \<And>x. x:A \<Longrightarrow> f(x):B(x); \<And>x. x:A \<Longrightarrow> B(x):\<box>\<rbrakk> |
|
115 |
\<Longrightarrow> Abs(A,f) : Prod(A,B)" |
|
45241 | 116 |
begin |
17252 | 117 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
118 |
lemmas [rules] = lam_sb pi_sb |
45241 | 119 |
|
120 |
end |
|
17252 | 121 |
|
17260 | 122 |
|
17252 | 123 |
locale LP2 = LP + L2 |
45241 | 124 |
begin |
17252 | 125 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
126 |
lemmas [rules] = lam_bs pi_bs lam_sb pi_sb |
45241 | 127 |
|
128 |
end |
|
17252 | 129 |
|
17260 | 130 |
|
17252 | 131 |
locale Lomega2 = L2 + Lomega |
45241 | 132 |
begin |
17252 | 133 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
134 |
lemmas [rules] = lam_bs pi_bs lam_bb pi_bb |
45241 | 135 |
|
136 |
end |
|
17252 | 137 |
|
17260 | 138 |
|
17252 | 139 |
locale LPomega = LP + Lomega |
45241 | 140 |
begin |
17252 | 141 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
142 |
lemmas [rules] = lam_bb pi_bb lam_sb pi_sb |
45241 | 143 |
|
144 |
end |
|
17252 | 145 |
|
17260 | 146 |
|
17252 | 147 |
locale CC = L2 + LP + Lomega |
45241 | 148 |
begin |
17252 | 149 |
|
49752
2bbb0013ff82
modernized dynamic "rules" -- avoid rebinding of static facts;
wenzelm
parents:
45242
diff
changeset
|
150 |
lemmas [rules] = lam_bs pi_bs lam_bb pi_bb lam_sb pi_sb |
17252 | 151 |
|
152 |
end |
|
45241 | 153 |
|
154 |
end |