aboutsummaryrefslogtreecommitdiff
path: root/Prod.thy
diff options
context:
space:
mode:
authorJosh Chen2018-09-16 11:03:48 +0200
committerJosh Chen2018-09-16 11:03:48 +0200
commitd4900ced2e071927d81a21a9127034941f258ec3 (patch)
treec0289b3fd8337a05baa7740ca3f5e84c57f539ca /Prod.thy
parent515872533295e8464799467303fff923b52a2c01 (diff)
parentf0999d07a0f41284ba84fae725a0186e0ec9ff5f (diff)
Reorganized HoTT_Base, updated theories
Diffstat (limited to 'Prod.thy')
-rw-r--r--Prod.thy42
1 files changed, 19 insertions, 23 deletions
diff --git a/Prod.thy b/Prod.thy
index 3b74531..a3cc347 100644
--- a/Prod.thy
+++ b/Prod.thy
@@ -12,14 +12,14 @@ begin
section \<open>Constants and syntax\<close>
axiomatization
- Prod :: "[Term, Typefam] \<Rightarrow> Term" and
- lambda :: "(Term \<Rightarrow> Term) \<Rightarrow> Term" (binder "\<^bold>\<lambda>" 30) and
- appl :: "[Term, Term] \<Rightarrow> Term" (infixl "`" 60)
+ Prod :: "[t, tf] \<Rightarrow> t" and
+ lambda :: "(t \<Rightarrow> t) \<Rightarrow> t" (binder "\<^bold>\<lambda>" 30) and
+ appl :: "[t, t] \<Rightarrow> t" (infixl "`" 60)
\<comment> \<open>Application binds tighter than abstraction.\<close>
syntax
- "_PROD" :: "[idt, Term, Term] \<Rightarrow> Term" ("(3\<Prod>_:_./ _)" 30)
- "_PROD_ASCII" :: "[idt, Term, Term] \<Rightarrow> Term" ("(3PROD _:_./ _)" 30)
+ "_PROD" :: "[idt, t, t] \<Rightarrow> t" ("(3\<Prod>_:_./ _)" 30)
+ "_PROD_ASCII" :: "[idt, t, t] \<Rightarrow> t" ("(3PROD _:_./ _)" 30)
text "The translations below bind the variable \<open>x\<close> in the expressions \<open>B\<close> and \<open>b\<close>."
@@ -29,24 +29,24 @@ translations
text "Nondependent functions are a special case."
-abbreviation Function :: "[Term, Term] \<Rightarrow> Term" (infixr "\<rightarrow>" 40)
+abbreviation Function :: "[t, t] \<Rightarrow> t" (infixr "\<rightarrow>" 40)
where "A \<rightarrow> B \<equiv> \<Prod>_: A. B"
section \<open>Type rules\<close>
axiomatization where
- Prod_form: "\<lbrakk>A: U(i); B: A \<longrightarrow> U(i)\<rbrakk> \<Longrightarrow> \<Prod>x:A. B(x): U(i)"
+ Prod_form: "\<lbrakk>A: U i; B: A \<longrightarrow> U i\<rbrakk> \<Longrightarrow> \<Prod>x:A. B x: U i"
and
- Prod_intro: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b(x): B(x); A: U(i)\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b(x): \<Prod>x:A. B(x)"
+ Prod_intro: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x: B x; A: U i\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b x: \<Prod>x:A. B x"
and
- Prod_elim: "\<lbrakk>f: \<Prod>x:A. B(x); a: A\<rbrakk> \<Longrightarrow> f`a: B(a)"
+ Prod_elim: "\<lbrakk>f: \<Prod>x:A. B x; a: A\<rbrakk> \<Longrightarrow> f`a: B a"
and
- Prod_appl: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b(x): B(x); a: A\<rbrakk> \<Longrightarrow> (\<^bold>\<lambda>x. b(x))`a \<equiv> b(a)"
+ Prod_appl: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x: B x; a: A\<rbrakk> \<Longrightarrow> (\<^bold>\<lambda>x. b x)`a \<equiv> b a"
and
- Prod_uniq: "f : \<Prod>x:A. B(x) \<Longrightarrow> \<^bold>\<lambda>x. (f`x) \<equiv> f"
+ Prod_uniq: "f : \<Prod>x:A. B x \<Longrightarrow> \<^bold>\<lambda>x. (f`x) \<equiv> f"
and
- Prod_eq: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b(x) \<equiv> b'(x); A: U(i)\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b(x) \<equiv> \<^bold>\<lambda>x. b'(x)"
+ Prod_eq: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x \<equiv> c x; A: U i\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b x \<equiv> \<^bold>\<lambda>x. c x"
text "
The Pure rules for \<open>\<equiv>\<close> only let us judge strict syntactic equality of object lambda expressions; Prod_eq is the actual definitional equality rule.
@@ -55,15 +55,15 @@ text "
"
text "
- In addition to the usual type rules, it is a meta-theorem that whenever \<open>\<Prod>x:A. B x: U(i)\<close> is derivable from some set of premises \<Gamma>, then so are \<open>A: U(i)\<close> and \<open>B: A \<longrightarrow> U(i)\<close>.
+ In addition to the usual type rules, it is a meta-theorem that whenever \<open>\<Prod>x:A. B x: U i\<close> is derivable from some set of premises \<Gamma>, then so are \<open>A: U i\<close> and \<open>B: A \<longrightarrow> U i\<close>.
That is to say, the following inference rules are admissible, and it simplifies proofs greatly to axiomatize them directly.
"
axiomatization where
- Prod_wellform1: "(\<Prod>x:A. B(x): U(i)) \<Longrightarrow> A: U(i)"
+ Prod_wellform1: "(\<Prod>x:A. B x: U i) \<Longrightarrow> A: U i"
and
- Prod_wellform2: "(\<Prod>x:A. B(x): U(i)) \<Longrightarrow> B: A \<longrightarrow> U(i)"
+ Prod_wellform2: "(\<Prod>x:A. B x: U i) \<Longrightarrow> B: A \<longrightarrow> U i"
text "Rule attribute declarations---set up various methods to use the type rules."
@@ -75,19 +75,15 @@ lemmas Prod_routine [intro] = Prod_form Prod_intro Prod_elim
section \<open>Function composition\<close>
-definition compose :: "[Term, Term] \<Rightarrow> Term" (infixr "o" 70) where "g o f \<equiv> \<^bold>\<lambda>x. g`(f`x)"
+definition compose :: "[t, t] \<Rightarrow> t" (infixr "o" 110) where "g o f \<equiv> \<^bold>\<lambda>x. g`(f`x)"
-syntax "_COMPOSE" :: "[Term, Term] \<Rightarrow> Term" (infixr "\<circ>" 70)
+syntax "_COMPOSE" :: "[t, t] \<Rightarrow> t" (infixr "\<circ>" 110)
translations "g \<circ> f" \<rightleftharpoons> "g o f"
-section \<open>Atomization\<close>
+section \<open>Polymorphic identity function\<close>
-text "
- Universal statements can be internalized within the theory; the following rule is admissible.
-" (* PROOF NEEDED *)
-
-axiomatization where Prod_atomize: "(\<^bold>\<lambda>x. b(x): \<Prod>x:A. B(x)) \<Longrightarrow> (\<And>x. x: A \<Longrightarrow> b(x): B(x))"
+abbreviation id :: t where "id \<equiv> \<^bold>\<lambda>x. x"
end