aboutsummaryrefslogtreecommitdiff
path: root/Prod.thy
blob: 42357933ddd263d6b4405efa1aaf1e34f9f8f22f (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
(********
Isabelle/HoTT: Dependent product (dependent function)
Feb 2019

********)

theory Prod
imports HoTT_Base HoTT_Methods

begin

section \<open>Basic type definitions\<close>

axiomatization
  Prod :: "[t, t \<Rightarrow> t] \<Rightarrow> t" and
  lam  :: "[t, t \<Rightarrow> t] \<Rightarrow> t" and
  app  :: "[t, t] \<Rightarrow> t"  ("(1_ ` _)" [120, 121] 120)
  \<comment> \<open>Application should bind tighter than abstraction.\<close>

syntax
  "_Prod"  :: "[idt, t, t] \<Rightarrow> t"  ("(3\<Prod>'(_: _')./ _)" 30)
  "_Prod'" :: "[idt, t, t] \<Rightarrow> t"  ("(3\<Prod>_: _./ _)" 30)
  "_lam"   :: "[idt, t, t] \<Rightarrow> t"  ("(3\<lambda>'(_: _')./ _)" 30)
  "_lam'"  :: "[idt, t, t] \<Rightarrow> t"  ("(3\<lambda>_: _./ _)" 30)
translations
  "\<Prod>(x: A). B" \<rightleftharpoons> "(CONST Prod) A (\<lambda>x. B)"
  "\<Prod>x: A. B" \<rightleftharpoons> "(CONST Prod) A (\<lambda>x. B)"
  "\<lambda>(x: A). b" \<rightleftharpoons> "(CONST lam) A (\<lambda>x. b)"
  "\<lambda>x: A. b" \<rightleftharpoons> "(CONST lam) A (\<lambda>x. b)"

text \<open>
The syntax translations above bind the variable @{term x} in the expressions @{term B} and @{term b}.
\<close>

text \<open>Non-dependent functions are a special case:\<close>

abbreviation Fun :: "[t, t] \<Rightarrow> t"  (infixr "\<rightarrow>" 40)
where "A \<rightarrow> B \<equiv> \<Prod>(_: A). B"

axiomatization where
\<comment> \<open>Type rules\<close>

  Prod_form: "\<lbrakk>A: U i; B: A \<leadsto> U i\<rbrakk> \<Longrightarrow> \<Prod>x: A. B x: U i" and

  Prod_intro: "\<lbrakk>A: U i; \<And>x. x: A \<Longrightarrow> b x: B x\<rbrakk> \<Longrightarrow> \<lambda>x: A. 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" and

  Prod_cmp: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x: B x; a: A\<rbrakk> \<Longrightarrow> (\<lambda>x: A. b x)`a \<equiv> b a" and

  Prod_uniq: "f: \<Prod>x: A. B x \<Longrightarrow> \<lambda>x: A. f`x \<equiv> f" and

\<comment> \<open>Congruence rules\<close>

  Prod_form_eq: "\<lbrakk>A: U i; B: A \<leadsto> U i; C: A \<leadsto> U i; \<And>x. x: A \<Longrightarrow> B x \<equiv> C x\<rbrakk>
    \<Longrightarrow> \<Prod>x: A. B x \<equiv> \<Prod>x: A. C x" and

  Prod_intro_eq: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x \<equiv> c x; A: U i\<rbrakk> \<Longrightarrow> \<lambda>x: A. b x \<equiv> \<lambda>x: A. c x"

text \<open>
The Pure rules for \<open>\<equiv>\<close> only let us judge strict syntactic equality of object lambda expressions.
The actual definitional equality rule in the type theory is @{thm Prod_intro_eq}.
Note that this is a separate rule from function extensionality.
\<close>

lemmas Prod_form [form]
lemmas Prod_routine [intro] = Prod_form Prod_intro Prod_elim
lemmas Prod_comp [comp] = Prod_cmp Prod_uniq
lemmas Prod_cong [cong] = Prod_form_eq Prod_intro_eq

section \<open>Function composition\<close>

definition compose :: "[t, t, t] \<Rightarrow> t"
where "compose A g f \<equiv> \<lambda>x: A. g`(f`x)"

declare compose_def [comp]

syntax "_compose" :: "[t, t] \<Rightarrow> t"  (infixr "o" 110)
parse_translation \<open>
let fun compose_tr ctxt [g, f] =
  let
    val [g, f] = [g, f] |> map (Util.prep_term ctxt)
    val dom =
      case f of
        Const ("Prod.lam", _) $ T $ _ => T
      | Const ("Prod.compose", _) $ T $ _ $ _ => T
      | _ =>
          case Typing.get_typing ctxt f of
            SOME (Const ("Prod.Prod", _) $ T $ _) => T
          | SOME t =>
              Exn.error ("Can't compose with a non-function " ^ Syntax.string_of_term ctxt f)
          | NONE =>
              Exn.error "Cannot infer domain of composition; please state this explicitly"
  in
    @{const compose} $ dom $ g $ f
  end
in
  [("_compose", compose_tr)]
end
\<close>

text \<open>Pretty-printing switch for composition; hides domain type information.\<close>

ML \<open>val pretty_compose = Attrib.setup_config_bool @{binding "pretty_compose"} (K true)\<close>

print_translation \<open>
let fun compose_tr' ctxt [A, g, f] =
  if Config.get ctxt pretty_compose
  then Syntax.const @{syntax_const "_compose"} $ g $ f
  else @{const compose} $ A $ g $ f
in
  [(@{const_syntax compose}, compose_tr')]
end
\<close>

lemma compose_assoc:
  assumes "A: U i" and "f: A \<rightarrow> B" and "g: B \<rightarrow> C" and "h: \<Prod>x: C. D x"
  shows "compose A (compose B h g) f \<equiv> compose A h (compose A g f)"
by (derive lems: assms cong)

lemma compose_comp:
  assumes "A: U i" and "\<And>x. x: A \<Longrightarrow> b x: B" and "\<And>x. x: B \<Longrightarrow> c x: C x"
  shows "(\<lambda>x: B. c x) o (\<lambda>x: A. b x) \<equiv> \<lambda>x: A. c (b x)"
by (derive lems: assms cong)

abbreviation id :: "t \<Rightarrow> t" where "id A \<equiv> \<lambda>x: A. x"

end